merged
authorpaulson
Sun, 10 Mar 2019 00:22:38 +0000
changeset 69881 6a6cdf34e980
parent 69879 2731278dfff9 (diff)
parent 69880 fe3c12990893 (current diff)
child 69893 1a7857abb75c
merged
--- a/src/HOL/String.thy	Sun Mar 10 00:09:45 2019 +0000
+++ b/src/HOL/String.thy	Sun Mar 10 00:22:38 2019 +0000
@@ -572,10 +572,28 @@
   is "map of_char"
   .
 
+qualified lemma asciis_of_zero [simp, code]:
+  "asciis_of_literal 0 = []"
+  by transfer simp
+
+qualified lemma asciis_of_Literal [simp, code]:
+  "asciis_of_literal (String.Literal b0 b1 b2 b3 b4 b5 b6 s) =
+    of_char (Char b0 b1 b2 b3 b4 b5 b6 False) # asciis_of_literal s "
+  by transfer simp
+
 qualified lift_definition literal_of_asciis :: "integer list \<Rightarrow> String.literal"
   is "map (String.ascii_of \<circ> char_of)"
   by auto
 
+qualified lemma literal_of_asciis_Nil [simp, code]:
+  "literal_of_asciis [] = 0"
+  by transfer simp
+
+qualified lemma literal_of_asciis_Cons [simp, code]:
+  "literal_of_asciis (k # ks) = (case char_of k
+    of Char b0 b1 b2 b3 b4 b5 b6 b7 \<Rightarrow> String.Literal b0 b1 b2 b3 b4 b5 b6 (literal_of_asciis ks))"
+  by (simp add: char_of_def) (transfer, simp add: char_of_def)
+
 qualified lemma literal_of_asciis_of_literal [simp]:
   "literal_of_asciis (asciis_of_literal s) = s"
 proof transfer
@@ -593,9 +611,14 @@
   "String.implode cs = literal_of_asciis (map of_char cs)"
   by transfer simp
 
-end
+qualified lemma equal_literal [code]:
+  "HOL.equal (String.Literal b0 b1 b2 b3 b4 b5 b6 s)
+    (String.Literal a0 a1 a2 a3 a4 a5 a6 r)
+    \<longleftrightarrow> (b0 \<longleftrightarrow> a0) \<and> (b1 \<longleftrightarrow> a1) \<and> (b2 \<longleftrightarrow> a2) \<and> (b3 \<longleftrightarrow> a3)
+      \<and> (b4 \<longleftrightarrow> a4) \<and> (b5 \<longleftrightarrow> a5) \<and> (b6 \<longleftrightarrow> a6) \<and> (s = r)"
+  by (simp add: equal)
 
-declare [[code drop: String.literal_of_asciis String.asciis_of_literal]]
+end
 
 
 subsubsection \<open>Technical code generation setup\<close>
@@ -619,7 +642,7 @@
 
 end
 
-code_reserved SML string Char
+code_reserved SML string String Char List
 code_reserved OCaml string String Char List
 code_reserved Haskell Prelude
 code_reserved Scala string
@@ -647,7 +670,7 @@
     and (Haskell) infixr 5 "++"
     and (Scala) infixl 7 "+"
 | constant String.literal_of_asciis \<rightharpoonup>
-    (SML) "!(String.implode/ o map (fn k => if 0 <= k andalso k < 128 then (Char.chr o IntInf.toInt) k else raise Fail \"Non-ASCII character in literal\"))"
+    (SML) "!(String.implode/ o List.map (fn k => if 0 <= k andalso k < 128 then (Char.chr o IntInf.toInt) k else raise Fail \"Non-ASCII character in literal\"))"
     and (OCaml) "!(let xs = _
       and chr k =
         let l = Big'_int.int'_of'_big'_int k
@@ -658,7 +681,7 @@
     and (Haskell) "map/ (let chr k | (0 <= k && k < 128) = Prelude.toEnum k :: Prelude.Char in chr . Prelude.fromInteger)"
     and (Scala) "\"\"/ ++/ _.map((k: BigInt) => if (BigInt(0) <= k && k < BigInt(128)) k.charValue else sys.error(\"Non-ASCII character in literal\"))"
 | constant String.asciis_of_literal \<rightharpoonup>
-    (SML) "!(map (fn c => let val k = Char.ord c in if k < 128 then IntInf.fromInt k else raise Fail \"Non-ASCII character in literal\" end) /o String.explode)"
+    (SML) "!(List.map (fn c => let val k = Char.ord c in if k < 128 then IntInf.fromInt k else raise Fail \"Non-ASCII character in literal\" end) /o String.explode)"
     and (OCaml) "!(let s = _ in let rec exp i l = if i < 0 then l else exp (i - 1) (let k = Char.code (String.get s i) in
       if k < 128 then Big'_int.big'_int'_of'_int k :: l else failwith \"Non-ASCII character in literal\") in exp (String.length s - 1) [])"
     and (Haskell) "map/ (let ord k | (k < 128) = Prelude.toInteger k in ord . (Prelude.fromEnum :: Prelude.Char -> Prelude.Int))"
--- a/src/Pure/Isar/outer_syntax.ML	Sun Mar 10 00:09:45 2019 +0000
+++ b/src/Pure/Isar/outer_syntax.ML	Sun Mar 10 00:22:38 2019 +0000
@@ -185,7 +185,7 @@
   Scan.ahead (before_command |-- Parse.position Parse.command) :|-- (fn (name, pos) =>
     let
       val keywords = Thy_Header.get_keywords thy;
-      val command_tags = Parse.command -- Parse.tags;
+      val command_tags = Parse.command -- Document_Source.tags;
       val tr0 =
         Toplevel.empty
         |> Toplevel.name name
--- a/src/Pure/Isar/parse.ML	Sun Mar 10 00:09:45 2019 +0000
+++ b/src/Pure/Isar/parse.ML	Sun Mar 10 00:22:38 2019 +0000
@@ -42,8 +42,6 @@
   val underscore: string parser
   val maybe: 'a parser -> 'a option parser
   val maybe_position: ('a * Position.T) parser -> ('a option * Position.T) parser
-  val tag_name: string parser
-  val tags: string list parser
   val opt_keyword: string -> bool parser
   val opt_bang: bool parser
   val begin: string parser
@@ -227,9 +225,6 @@
 val int = Scan.optional (minus >> K ~1) 1 -- nat >> op *;
 val real = float_number >> Value.parse_real || int >> Real.fromInt;
 
-val tag_name = group (fn () => "tag name") (short_ident || string);
-val tags = Scan.repeat ($$$ "%" |-- !!! tag_name);
-
 fun opt_keyword s = Scan.optional ($$$ "(" |-- !!! (($$$ s >> K true) --| $$$ ")")) false;
 val opt_bang = Scan.optional ($$$ "!" >> K true) false;
 
--- a/src/Pure/Isar/toplevel.ML	Sun Mar 10 00:09:45 2019 +0000
+++ b/src/Pure/Isar/toplevel.ML	Sun Mar 10 00:22:38 2019 +0000
@@ -33,9 +33,11 @@
   val empty: transition
   val name_of: transition -> string
   val pos_of: transition -> Position.T
+  val timing_of: transition -> Time.time
   val type_error: transition -> string
   val name: string -> transition -> transition
   val position: Position.T -> transition -> transition
+  val timing: Time.time -> transition -> transition
   val init_theory: (unit -> theory) -> transition -> transition
   val is_init: transition -> bool
   val modify_init: (unit -> theory) -> transition -> transition
@@ -77,8 +79,6 @@
   val exec_id: Document_ID.exec -> transition -> transition
   val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
   val add_hook: (transition -> state -> state -> unit) -> unit
-  val get_timing: transition -> Time.time
-  val put_timing: Time.time -> transition -> transition
   val transition: bool -> transition -> state -> state * (exn * string) option
   val command_errors: bool -> transition -> state -> Runtime.error list * state option
   val command_exception: bool -> transition -> state -> state
@@ -101,27 +101,30 @@
 (* datatype node *)
 
 datatype node =
-  Theory of generic_theory * Proof.context option
-    (*theory with presentation context*) |
+  Theory of generic_theory
+    (*global or local theory*) |
   Proof of Proof_Node.T * ((Proof.context -> generic_theory) * generic_theory)
     (*proof node, finish, original theory*) |
   Skipped_Proof of int * (generic_theory * generic_theory);
     (*proof depth, resulting theory, original theory*)
 
-val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE;
+val theory_node = fn Theory gthy => SOME gthy | _ => NONE;
 val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE;
 val skipped_proof_node = fn Skipped_Proof _ => true | _ => false;
 
-fun cases_node f _ (Theory (gthy, _)) = f gthy
+fun cases_node f _ (Theory gthy) = f gthy
   | cases_node _ g (Proof (prf, _)) = g (Proof_Node.current prf)
   | cases_node f _ (Skipped_Proof (_, (gthy, _))) = f gthy;
 
 
 (* datatype state *)
 
-datatype state = State of node option * node option;  (*current, previous*)
+type node_presentation = node * Proof.context;  (*node with presentation context*)
+fun node_presentation0 node = (node, cases_node Context.proof_of Proof.context_of node);
 
-fun theory_toplevel thy = State (SOME (Theory (Context.Theory thy, NONE)), NONE);
+datatype state = State of node_presentation option * node_presentation option; (*current, previous*)
+
+fun theory_toplevel thy = State (SOME (node_presentation0 (Theory (Context.Theory thy))), NONE);
 
 val toplevel = State (NONE, NONE);
 
@@ -129,23 +132,23 @@
   | is_toplevel _ = false;
 
 fun level (State (NONE, _)) = 0
-  | level (State (SOME (Theory _), _)) = 0
-  | level (State (SOME (Proof (prf, _)), _)) = Proof.level (Proof_Node.current prf)
-  | level (State (SOME (Skipped_Proof (d, _)), _)) = d + 1;   (*different notion of proof depth!*)
+  | level (State (SOME (Theory _, _), _)) = 0
+  | level (State (SOME (Proof (prf, _), _), _)) = Proof.level (Proof_Node.current prf)
+  | level (State (SOME (Skipped_Proof (d, _), _), _)) = d + 1;   (*different notion of proof depth!*)
 
-fun str_of_state (State (NONE, SOME (Theory (Context.Theory thy, _)))) =
+fun str_of_state (State (NONE, SOME (Theory (Context.Theory thy), _))) =
       "at top level, result theory " ^ quote (Context.theory_name thy)
   | str_of_state (State (NONE, _)) = "at top level"
-  | str_of_state (State (SOME (Theory (Context.Theory _, _)), _)) = "in theory mode"
-  | str_of_state (State (SOME (Theory (Context.Proof _, _)), _)) = "in local theory mode"
-  | str_of_state (State (SOME (Proof _), _)) = "in proof mode"
-  | str_of_state (State (SOME (Skipped_Proof _), _)) = "in skipped proof mode";
+  | str_of_state (State (SOME (Theory (Context.Theory _), _), _)) = "in theory mode"
+  | str_of_state (State (SOME (Theory (Context.Proof _), _), _)) = "in local theory mode"
+  | str_of_state (State (SOME (Proof _, _), _)) = "in proof mode"
+  | str_of_state (State (SOME (Skipped_Proof _, _), _)) = "in skipped proof mode";
 
 
 (* current node *)
 
 fun node_of (State (NONE, _)) = raise UNDEF
-  | node_of (State (SOME node, _)) = node;
+  | node_of (State (SOME (node, _), _)) = node;
 
 fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state));
 fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state));
@@ -154,7 +157,7 @@
 fun node_case f g state = cases_node f g (node_of state);
 
 fun previous_theory_of (State (_, NONE)) = NONE
-  | previous_theory_of (State (_, SOME prev)) =
+  | previous_theory_of (State (_, SOME (prev, _))) =
       SOME (cases_node Context.theory_of Proof.theory_of prev);
 
 val context_of = node_case Context.proof_of Proof.context_of;
@@ -167,10 +170,10 @@
     Proof (prf, _) => Proof_Node.position prf
   | _ => ~1);
 
-fun is_end_theory (State (NONE, SOME (Theory (Context.Theory _, _)))) = true
+fun is_end_theory (State (NONE, SOME (Theory (Context.Theory _), _))) = true
   | is_end_theory _ = false;
 
-fun end_theory _ (State (NONE, SOME (Theory (Context.Theory thy, _)))) = thy
+fun end_theory _ (State (NONE, SOME (Theory (Context.Theory thy), _))) = thy
   | end_theory pos _ = error ("Malformed theory" ^ Position.here pos);
 
 
@@ -183,10 +186,9 @@
 );
 
 fun presentation_context0 state =
-  (case try node_of state of
-    SOME (Theory (_, SOME ctxt)) => ctxt
-  | SOME node => cases_node Context.proof_of Proof.context_of node
-  | NONE =>
+  (case state of
+    State (SOME (_, ctxt), _) => ctxt
+  | State (NONE, _) =>
       (case try Theory.get_pure () of
         SOME thy => Proof_Context.init_global thy
       | NONE => raise UNDEF));
@@ -197,7 +199,7 @@
 
 fun presentation_state ctxt =
   (case Presentation_State.get ctxt of
-    NONE => State (SOME (Theory (Context.Proof ctxt, SOME ctxt)), NONE)
+    NONE => State (SOME (node_presentation0 (Theory (Context.Proof ctxt))), NONE)
   | SOME state => state);
 
 
@@ -210,7 +212,7 @@
       let
         val gthy =
           (case node of
-            Theory (gthy, _) => gthy
+            Theory gthy => gthy
           | Proof (_, (_, gthy)) => gthy
           | Skipped_Proof (_, (_, gthy)) => gthy);
         val lthy = Context.cases Named_Target.theory_init I gthy;
@@ -237,24 +239,17 @@
 
 exception FAILURE of state * exn;
 
-local
-
-fun reset_presentation (Theory (gthy, _)) = Theory (gthy, NONE)
-  | reset_presentation node = node;
-
-in
-
 fun apply_transaction f g node =
   let
-    val cont_node = reset_presentation node;
-    val context = cases_node I (Context.Proof o Proof.context_of) cont_node;
-    fun state_error e nd = (State (SOME nd, SOME cont_node), e);
+    val node_pr = node_presentation0 node;
+    val context = cases_node I (Context.Proof o Proof.context_of) node;
+    fun state_error e node_pr' = (State (SOME node_pr', SOME node_pr), e);
 
     val (result, err) =
-      cont_node
+      node
       |> Runtime.controlled_execution (SOME context) f
       |> state_error NONE
-      handle exn => state_error (SOME exn) cont_node;
+      handle exn => state_error (SOME exn) node_pr;
   in
     (case err of
       NONE => tap g result
@@ -263,30 +258,34 @@
 
 val exit_transaction =
   apply_transaction
-    (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (Theory.end_theory thy), NONE)
-      | node => node) (K ())
-  #> (fn State (node', _) => State (NONE, node'));
-
-end;
+    ((fn Theory (Context.Theory thy) => Theory (Context.Theory (Theory.end_theory thy))
+       | node => node) #> node_presentation0)
+    (K ())
+  #> (fn State (node_pr', _) => State (NONE, node_pr'));
 
 
 (* primitive transitions *)
 
 datatype trans =
-  Init of unit -> theory |               (*init theory*)
-  Exit |                                 (*formal exit of theory*)
-  Keep of bool -> state -> unit |        (*peek at state*)
-  Transaction of (bool -> node -> node) * (state -> unit);  (*node transaction and presentation*)
+  (*init theory*)
+  Init of unit -> theory |
+  (*formal exit of theory*)
+  Exit |
+  (*peek at state*)
+  Keep of bool -> state -> unit |
+  (*node transaction and presentation*)
+  Transaction of (bool -> node -> node_presentation) * (state -> unit);
 
 local
 
 fun apply_tr _ (Init f) (State (NONE, _)) =
-      State (SOME (Theory (Context.Theory (Runtime.controlled_execution NONE f ()), NONE)), NONE)
-  | apply_tr _ Exit (State (SOME (state as Theory (Context.Theory _, _)), _)) =
-      exit_transaction state
+      let val node = Theory (Context.Theory (Runtime.controlled_execution NONE f ()))
+      in State (SOME (node_presentation0 node), NONE) end
+  | apply_tr _ Exit (State (SOME (node as Theory (Context.Theory _), _), _)) =
+      exit_transaction node
   | apply_tr int (Keep f) state =
       Runtime.controlled_execution (try generic_theory_of state) (fn x => tap (f int) x) state
-  | apply_tr int (Transaction (f, g)) (State (SOME node, _)) =
+  | apply_tr int (Transaction (f, g)) (State (SOME (node, _), _)) =
       apply_transaction (fn x => f int x) g node
   | apply_tr _ _ _ = raise UNDEF;
 
@@ -327,6 +326,7 @@
 
 fun name_of (Transition {name, ...}) = name;
 fun pos_of (Transition {pos, ...}) = pos;
+fun timing_of (Transition {timing, ...}) = timing;
 
 fun command_msg msg tr =
   msg ^ "command " ^ quote (Markup.markup Markup.keyword1 (name_of tr)) ^
@@ -344,6 +344,9 @@
 fun position pos = map_transition (fn (name, _, timing, trans) =>
   (name, pos, timing, trans));
 
+fun timing timing = map_transition (fn (name, pos, _, trans) =>
+  (name, pos, timing, trans));
+
 fun add_trans tr = map_transition (fn (name, pos, timing, trans) =>
   (name, pos, timing, tr :: trans));
 
@@ -365,6 +368,7 @@
 
 fun present_transaction f g = add_trans (Transaction (f, g));
 fun transaction f = present_transaction f (K ());
+fun transaction0 f = present_transaction (node_presentation0 oo f) (K ());
 
 fun keep f = add_trans (Keep (fn _ => f));
 
@@ -384,22 +388,22 @@
 (* theory transitions *)
 
 fun generic_theory f = transaction (fn _ =>
-  (fn Theory (gthy, _) => Theory (f gthy, NONE)
+  (fn Theory gthy => node_presentation0 (Theory (f gthy))
     | _ => raise UNDEF));
 
 fun theory' f = transaction (fn int =>
-  (fn Theory (Context.Theory thy, _) =>
+  (fn Theory (Context.Theory thy) =>
       let val thy' = thy
         |> Sign.new_group
         |> f int
         |> Sign.reset_group;
-      in Theory (Context.Theory thy', NONE) end
+      in node_presentation0 (Theory (Context.Theory thy')) end
     | _ => raise UNDEF));
 
 fun theory f = theory' (K f);
 
 fun begin_local_theory begin f = transaction (fn _ =>
-  (fn Theory (Context.Theory thy, _) =>
+  (fn Theory (Context.Theory thy) =>
         let
           val lthy = f thy;
           val gthy = if begin then Context.Proof lthy else Context.Theory (Named_Target.exit lthy);
@@ -407,21 +411,21 @@
             (case Local_Theory.pretty lthy of
               [] => ()
             | prts => Output.state (Pretty.string_of (Pretty.chunks prts)));
-        in Theory (gthy, SOME lthy) end
+        in (Theory gthy, lthy) end
     | _ => raise UNDEF));
 
 val end_local_theory = transaction (fn _ =>
-  (fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (Named_Target.exit lthy), SOME lthy)
+  (fn Theory (Context.Proof lthy) => (Theory (Context.Theory (Named_Target.exit lthy)), lthy)
     | _ => raise UNDEF));
 
-fun open_target f = transaction (fn _ =>
-  (fn Theory (gthy, _) =>
+fun open_target f = transaction0 (fn _ =>
+  (fn Theory gthy =>
         let val lthy = f gthy
-        in Theory (Context.Proof lthy, SOME lthy) end
+        in Theory (Context.Proof lthy) end
     | _ => raise UNDEF));
 
 val close_target = transaction (fn _ =>
-  (fn Theory (Context.Proof lthy, _) =>
+  (fn Theory (Context.Proof lthy) =>
         (case try Local_Theory.close_target lthy of
           SOME ctxt' =>
             let
@@ -429,7 +433,7 @@
                 if can Local_Theory.assert ctxt'
                 then Context.Proof ctxt'
                 else Context.Theory (Proof_Context.theory_of ctxt');
-            in Theory (gthy', SOME lthy) end
+            in (Theory gthy', lthy) end
         | NONE => raise UNDEF)
     | _ => raise UNDEF));
 
@@ -438,7 +442,7 @@
   | restricted_context NONE = I;
 
 fun local_theory' restricted target f = present_transaction (fn int =>
-  (fn Theory (gthy, _) =>
+  (fn Theory gthy =>
         let
           val (finish, lthy) = Named_Target.switch target gthy;
           val lthy' = lthy
@@ -446,16 +450,16 @@
             |> Local_Theory.new_group
             |> f int
             |> Local_Theory.reset_group;
-        in Theory (finish lthy', SOME lthy') end
+        in (Theory (finish lthy'), lthy') end
     | _ => raise UNDEF))
   (K ());
 
 fun local_theory restricted target f = local_theory' restricted target (K f);
 
 fun present_local_theory target = present_transaction (fn _ =>
-  (fn Theory (gthy, _) =>
+  (fn Theory gthy =>
         let val (finish, lthy) = Named_Target.switch target gthy;
-        in Theory (finish lthy, SOME lthy) end
+        in (Theory (finish lthy), lthy) end
     | _ => raise UNDEF));
 
 
@@ -468,16 +472,16 @@
             let
               val ctxt' = f int state;
               val gthy' = finish ctxt';
-            in Theory (gthy', SOME ctxt') end
+            in (Theory gthy', ctxt') end
           else raise UNDEF
         end
-    | Skipped_Proof (0, (gthy, _)) => Theory (gthy, NONE)
+    | Skipped_Proof (0, (gthy, _)) => node_presentation0 (Theory gthy)
     | _ => raise UNDEF));
 
 local
 
-fun begin_proof init = transaction (fn int =>
-  (fn Theory (gthy, _) =>
+fun begin_proof init = transaction0 (fn int =>
+  (fn Theory gthy =>
     let
       val (finish, prf) = init int gthy;
       val document = Options.default_string "document";
@@ -518,14 +522,14 @@
 
 end;
 
-val forget_proof = transaction (fn _ =>
+val forget_proof = transaction0 (fn _ =>
   (fn Proof (prf, (_, orig_gthy)) =>
         if Proof.is_notepad (Proof_Node.current prf) then raise UNDEF
-        else Theory (orig_gthy, NONE)
-    | Skipped_Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
+        else Theory orig_gthy
+    | Skipped_Proof (_, (_, orig_gthy)) => Theory orig_gthy
     | _ => raise UNDEF));
 
-fun proofs' f = transaction (fn int =>
+fun proofs' f = transaction0 (fn int =>
   (fn Proof (prf, x) => Proof (Proof_Node.applys (f int) prf, x)
     | skip as Skipped_Proof _ => skip
     | _ => raise UNDEF));
@@ -537,20 +541,20 @@
 
 (* skipped proofs *)
 
-fun actual_proof f = transaction (fn _ =>
+fun actual_proof f = transaction0 (fn _ =>
   (fn Proof (prf, x) => Proof (f prf, x)
     | _ => raise UNDEF));
 
-fun skip_proof f = transaction (fn _ =>
+fun skip_proof f = transaction0 (fn _ =>
   (fn skip as Skipped_Proof _ => (f (); skip)
     | _ => raise UNDEF));
 
-val skip_proof_open = transaction (fn _ =>
+val skip_proof_open = transaction0 (fn _ =>
   (fn Skipped_Proof (d, x) => Skipped_Proof (d + 1, x)
     | _ => raise UNDEF));
 
-val skip_proof_close = transaction (fn _ =>
-  (fn Skipped_Proof (0, (gthy, _)) => Theory (gthy, NONE)
+val skip_proof_close = transaction0 (fn _ =>
+  (fn Skipped_Proof (0, (gthy, _)) => Theory gthy
     | Skipped_Proof (d, x) => Skipped_Proof (d - 1, x)
     | _ => raise UNDEF));
 
@@ -582,9 +586,6 @@
 
 (* apply transitions *)
 
-fun get_timing (Transition {timing, ...}) = timing;
-fun put_timing timing = map_transition (fn (name, pos, _, trans) => (name, pos, timing, trans));
-
 local
 
 fun app int (tr as Transition {trans, ...}) =
@@ -639,8 +640,8 @@
 
 val reset_proof =
   reset_state is_proof
-    (transaction (fn _ =>
-      (fn Theory (gthy, _) => Skipped_Proof (0, (gthy, gthy))
+    (transaction0 (fn _ =>
+      (fn Theory gthy => Skipped_Proof (0, (gthy, gthy))
         | _ => raise UNDEF)));
 
 val reset_notepad =
@@ -678,7 +679,7 @@
 
 fun timing_estimate elem =
   let val trs = tl (Thy_Element.flat_element elem)
-  in fold (fn tr => fn t => get_timing tr + t) trs Time.zeroTime end;
+  in fold (fn tr => fn t => timing_of tr + t) trs Time.zeroTime end;
 
 fun future_proofs_enabled estimate st =
   (case try proof_of st of
@@ -724,10 +725,10 @@
                   {name = "Toplevel.future_proof", pos = pos_of head_tr, pri = ~1}
                   (fn () =>
                     let
-                      val State (SOME (Proof (prf, (_, orig_gthy))), prev) = st';
-                      val prf' = Proof_Node.apply (K state) prf;
+                      val State (SOME (Proof (prf, (_, orig_gthy)), _), prev) = st';
+                      val node' = Proof (Proof_Node.apply (K state) prf, (finish, orig_gthy));
                       val (result, result_state) =
-                        State (SOME (Proof (prf', (finish, orig_gthy))), prev)
+                        State (SOME (node_presentation0 node'), prev)
                         |> fold_map (element_result keywords) body_elems ||> command end_tr;
                     in (Result_List result, presentation_context0 result_state) end))
               #> (fn (res, state') => state' |> put_result (Result_Future res));
--- a/src/Pure/ROOT.ML	Sun Mar 10 00:09:45 2019 +0000
+++ b/src/Pure/ROOT.ML	Sun Mar 10 00:22:38 2019 +0000
@@ -206,6 +206,7 @@
 ML_file "Isar/keyword.ML";
 ML_file "Isar/token.ML";
 ML_file "Isar/parse.ML";
+ML_file "Thy/document_source.ML";
 ML_file "Thy/thy_header.ML";
 
 (*proof context*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Thy/document_source.ML	Sun Mar 10 00:22:38 2019 +0000
@@ -0,0 +1,51 @@
+(*  Title:      Pure/Thy/document_source.ML
+    Author:     Makarius
+
+Document source for presentation.
+*)
+
+signature DOCUMENT_SOURCE =
+sig
+  val is_white: Token.T -> bool
+  val is_black: Token.T -> bool
+  val is_white_comment: Token.T -> bool
+  val is_black_comment: Token.T -> bool
+  val is_improper: Token.T -> bool
+  val improper: Token.T list parser
+  val improper_end: Token.T list parser
+  val blank_end: Token.T list parser
+  val tag: string parser
+  val tags: string list parser
+end;
+
+structure Document_Source: DOCUMENT_SOURCE =
+struct
+
+(* white space and comments *)
+
+(*NB: arranging white space around command spans is a black art*)
+
+val is_white = Token.is_space orf Token.is_informal_comment;
+val is_black = not o is_white;
+
+val is_white_comment = Token.is_informal_comment;
+val is_black_comment = Token.is_formal_comment;
+
+
+val space_proper =
+  Scan.one Token.is_blank -- Scan.many is_white_comment -- Scan.one is_black;
+
+val is_improper = not o (is_black orf Token.is_begin_ignore orf Token.is_end_ignore);
+val improper = Scan.many is_improper;
+val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper));
+val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one Token.is_blank));
+
+
+(* tags *)
+
+val tag_name = Parse.group (fn () => "document tag name") (Parse.short_ident || Parse.string);
+
+val tag = (improper -- Parse.$$$ "%" -- improper) |-- Parse.!!! (tag_name --| blank_end);
+val tags = Scan.repeat tag;
+
+end;
--- a/src/Pure/Thy/thy_header.ML	Sun Mar 10 00:09:45 2019 +0000
+++ b/src/Pure/Thy/thy_header.ML	Sun Mar 10 00:22:38 2019 +0000
@@ -137,7 +137,7 @@
 
 val keyword_spec =
   Parse.group (fn () => "outer syntax keyword specification")
-    (Parse.name -- opt_files -- Parse.tags);
+    (Parse.name -- opt_files -- Document_Source.tags);
 
 val keyword_decl =
   Scan.repeat1 Parse.string_position --
@@ -175,10 +175,10 @@
     Parse.command_name textN ||
     Parse.command_name txtN ||
     Parse.command_name text_rawN) --
-  Parse.tags -- Parse.!!! Parse.document_source;
+  Document_Source.tags -- Parse.!!! Parse.document_source;
 
 val parse_header =
-  (Scan.repeat heading -- Parse.command_name theoryN -- Parse.tags) |-- Parse.!!! args;
+  (Scan.repeat heading -- Parse.command_name theoryN -- Document_Source.tags) |-- Parse.!!! args;
 
 fun read_tokens pos toks =
   filter Token.is_proper toks
--- a/src/Pure/Thy/thy_info.ML	Sun Mar 10 00:09:45 2019 +0000
+++ b/src/Pure/Thy/thy_info.ML	Sun Mar 10 00:22:38 2019 +0000
@@ -285,7 +285,7 @@
     fun prepare_span st span =
       Command_Span.content span
       |> Command.read keywords (Command.read_thy st) master_dir init ([], ~1)
-      |> (fn tr => Toplevel.put_timing (last_timing tr) tr);
+      |> (fn tr => Toplevel.timing (last_timing tr) tr);
 
     fun element_result span_elem (st, _) =
       let
--- a/src/Pure/Thy/thy_output.ML	Sun Mar 10 00:09:45 2019 +0000
+++ b/src/Pure/Thy/thy_output.ML	Sun Mar 10 00:22:38 2019 +0000
@@ -170,15 +170,6 @@
 
 (** present theory source **)
 
-(*NB: arranging white space around command spans is a black art*)
-
-val is_white = Token.is_space orf Token.is_informal_comment;
-val is_black = not o is_white;
-
-val is_white_comment = Token.is_informal_comment;
-val is_black_comment = Token.is_formal_comment;
-
-
 (* presentation tokens *)
 
 datatype token =
@@ -191,8 +182,8 @@
 fun basic_token pred (Basic_Token tok) = pred tok
   | basic_token _ _ = false;
 
-val white_token = basic_token is_white;
-val white_comment_token = basic_token is_white_comment;
+val white_token = basic_token Document_Source.is_white;
+val white_comment_token = basic_token Document_Source.is_white_comment;
 val blank_token = basic_token Token.is_blank;
 val newline_token = basic_token Token.is_newline;
 
@@ -348,14 +339,6 @@
 val markup_true = "\\isamarkuptrue%\n";
 val markup_false = "\\isamarkupfalse%\n";
 
-val space_proper =
-  Scan.one Token.is_blank -- Scan.many is_white_comment -- Scan.one is_black;
-
-val is_improper = not o (is_black orf Token.is_begin_ignore orf Token.is_end_ignore);
-val improper = Scan.many is_improper;
-val improper_end = Scan.repeat (Scan.unless space_proper (Scan.one is_improper));
-val blank_end = Scan.repeat (Scan.unless space_proper (Scan.one Token.is_blank));
-
 val opt_newline = Scan.option (Scan.one Token.is_newline);
 
 val ignore =
@@ -365,11 +348,10 @@
     (if d = 0 then Scan.fail_with (K (fn () => "Bad nesting of meta-comments")) else opt_newline)
     >> pair (d - 1));
 
-val tag = (improper -- Parse.$$$ "%" -- improper) |-- Parse.!!! (Parse.tag_name --| blank_end);
-
 val locale =
-  Scan.option ((Parse.$$$ "(" -- improper -- Parse.$$$ "in") |--
-    Parse.!!! (improper |-- Parse.name --| (improper -- Parse.$$$ ")")));
+  Scan.option ((Parse.$$$ "(" -- Document_Source.improper -- Parse.$$$ "in") |--
+    Parse.!!!
+      (Document_Source.improper |-- Parse.name --| (Document_Source.improper -- Parse.$$$ ")")));
 
 in
 
@@ -386,25 +368,26 @@
       >> (fn d => (NONE, (Ignore_Token, ("", d))));
 
     fun markup pred mk flag = Scan.peek (fn d =>
-      improper |--
+      Document_Source.improper |--
         Parse.position (Scan.one (fn tok =>
           Token.is_command tok andalso pred keywords (Token.content_of tok))) --
-      Scan.repeat tag --
-      Parse.!!!! ((improper -- locale -- improper) |-- Parse.document_source --| improper_end)
-      >> (fn (((tok, pos'), tags), source) =>
-        let val name = Token.content_of tok
-        in (SOME (name, pos', tags), (mk (name, source), (flag, d))) end));
+      Scan.repeat Document_Source.tag --
+      Parse.!!!! ((Document_Source.improper -- locale -- Document_Source.improper) |--
+        Parse.document_source --| Document_Source.improper_end)
+          >> (fn (((tok, pos'), tags), source) =>
+            let val name = Token.content_of tok
+            in (SOME (name, pos', tags), (mk (name, source), (flag, d))) end));
 
     val command = Scan.peek (fn d =>
-      Scan.optional (Scan.one Token.is_command_modifier ::: improper) [] --
-      Scan.one Token.is_command -- Scan.repeat tag
+      Scan.optional (Scan.one Token.is_command_modifier ::: Document_Source.improper) [] --
+      Scan.one Token.is_command -- Document_Source.tags
       >> (fn ((cmd_mod, cmd), tags) =>
         map (fn tok => (NONE, (Basic_Token tok, ("", d)))) cmd_mod @
           [(SOME (Token.content_of cmd, Token.pos_of cmd, tags),
             (Basic_Token cmd, (markup_false, d)))]));
 
     val cmt = Scan.peek (fn d =>
-      Scan.one is_black_comment >> (fn tok => (NONE, (Basic_Token tok, ("", d)))));
+      Scan.one Document_Source.is_black_comment >> (fn tok => (NONE, (Basic_Token tok, ("", d)))));
 
     val other = Scan.peek (fn d =>
        Parse.not_eof >> (fn tok => (NONE, (Basic_Token tok, ("", d)))));