clarified document preparation vs. skip_proofs;
1 (* Title: Pure/Isar/toplevel.ML
2 Author: Markus Wenzel, TU Muenchen
4 Isabelle/Isar toplevel transactions.
11 val theory_toplevel: theory -> state
13 val is_toplevel: state -> bool
14 val is_theory: state -> bool
15 val is_proof: state -> bool
16 val is_skipped_proof: state -> bool
17 val level: state -> int
18 val presentation_context_of: state -> Proof.context
19 val previous_context_of: state -> Proof.context option
20 val context_of: state -> Proof.context
21 val generic_theory_of: state -> generic_theory
22 val theory_of: state -> theory
23 val proof_of: state -> Proof.state
24 val proof_position_of: state -> int
25 val end_theory: Position.T -> state -> theory
26 val pretty_context: state -> Pretty.T list
27 val pretty_state: state -> Pretty.T list
28 val string_of_state: state -> string
29 val pretty_abstract: state -> Pretty.T
32 val name_of: transition -> string
33 val pos_of: transition -> Position.T
34 val type_error: transition -> string
35 val name: string -> transition -> transition
36 val position: Position.T -> transition -> transition
37 val init_theory: (unit -> theory) -> transition -> transition
38 val is_init: transition -> bool
39 val modify_init: (unit -> theory) -> transition -> transition
40 val exit: transition -> transition
41 val keep: (state -> unit) -> transition -> transition
42 val keep': (bool -> state -> unit) -> transition -> transition
43 val keep_proof: (state -> unit) -> transition -> transition
44 val ignored: Position.T -> transition
45 val is_ignored: transition -> bool
46 val malformed: Position.T -> string -> transition
47 val generic_theory: (generic_theory -> generic_theory) -> transition -> transition
48 val theory': (bool -> theory -> theory) -> transition -> transition
49 val theory: (theory -> theory) -> transition -> transition
50 val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition
51 val end_local_theory: transition -> transition
52 val open_target: (generic_theory -> local_theory) -> transition -> transition
53 val close_target: transition -> transition
54 val local_theory': (bool * Position.T) option -> (xstring * Position.T) option ->
55 (bool -> local_theory -> local_theory) -> transition -> transition
56 val local_theory: (bool * Position.T) option -> (xstring * Position.T) option ->
57 (local_theory -> local_theory) -> transition -> transition
58 val present_local_theory: (xstring * Position.T) option -> (state -> unit) ->
59 transition -> transition
60 val local_theory_to_proof': (bool * Position.T) option -> (xstring * Position.T) option ->
61 (bool -> local_theory -> Proof.state) -> transition -> transition
62 val local_theory_to_proof: (bool * Position.T) option -> (xstring * Position.T) option ->
63 (local_theory -> Proof.state) -> transition -> transition
64 val theory_to_proof: (theory -> Proof.state) -> transition -> transition
65 val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition
66 val forget_proof: bool -> transition -> transition
67 val proofs': (bool -> Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition
68 val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition
69 val proofs: (Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition
70 val proof: (Proof.state -> Proof.state) -> transition -> transition
71 val actual_proof: (Proof_Node.T -> Proof_Node.T) -> transition -> transition
72 val skip_proof: (unit -> unit) -> transition -> transition
73 val skip_proof_open: transition -> transition
74 val skip_proof_close: transition -> transition
75 val exec_id: Document_ID.exec -> transition -> transition
76 val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
77 val add_hook: (transition -> state -> state -> unit) -> unit
78 val get_timing: transition -> Time.time
79 val put_timing: Time.time -> transition -> transition
80 val transition: bool -> transition -> state -> state * (exn * string) option
81 val command_errors: bool -> transition -> state -> Runtime.error list * state option
82 val command_exception: bool -> transition -> state -> state
83 val reset_theory: state -> state option
84 val reset_proof: state -> state option
86 val join_results: result -> (transition * state) list
87 val element_result: Keyword.keywords -> transition Thy_Syntax.element -> state -> result * state
90 structure Toplevel: TOPLEVEL =
93 (** toplevel state **)
95 exception UNDEF = Runtime.UNDEF;
101 Theory of generic_theory * Proof.context option
102 (*theory with presentation context*) |
103 Proof of Proof_Node.T * ((Proof.context -> generic_theory) * generic_theory)
104 (*proof node, finish, original theory*) |
105 Skipped_Proof of int * (generic_theory * generic_theory);
106 (*proof depth, resulting theory, original theory*)
108 val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE;
109 val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE;
110 val skipped_proof_node = fn Skipped_Proof _ => true | _ => false;
112 fun cases_node f _ (Theory (gthy, _)) = f gthy
113 | cases_node _ g (Proof (prf, _)) = g (Proof_Node.current prf)
114 | cases_node f _ (Skipped_Proof (_, (gthy, _))) = f gthy;
116 val context_node = cases_node Context.proof_of Proof.context_of;
121 datatype state = State of node option * node option; (*current, previous*)
123 fun theory_toplevel thy = State (SOME (Theory (Context.Theory thy, NONE)), NONE);
125 val toplevel = State (NONE, NONE);
127 fun is_toplevel (State (NONE, _)) = true
128 | is_toplevel _ = false;
130 fun level (State (NONE, _)) = 0
131 | level (State (SOME (Theory _), _)) = 0
132 | level (State (SOME (Proof (prf, _)), _)) = Proof.level (Proof_Node.current prf)
133 | level (State (SOME (Skipped_Proof (d, _)), _)) = d + 1; (*different notion of proof depth!*)
135 fun str_of_state (State (NONE, SOME (Theory (Context.Theory thy, _)))) =
136 "at top level, result theory " ^ quote (Context.theory_name thy)
137 | str_of_state (State (NONE, _)) = "at top level"
138 | str_of_state (State (SOME (Theory (Context.Theory _, _)), _)) = "in theory mode"
139 | str_of_state (State (SOME (Theory (Context.Proof _, _)), _)) = "in local theory mode"
140 | str_of_state (State (SOME (Proof _), _)) = "in proof mode"
141 | str_of_state (State (SOME (Skipped_Proof _), _)) = "in skipped proof mode";
146 fun node_of (State (NONE, _)) = raise UNDEF
147 | node_of (State (SOME node, _)) = node;
149 fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state));
150 fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state));
151 fun is_skipped_proof state = not (is_toplevel state) andalso skipped_proof_node (node_of state);
153 fun node_case f g state = cases_node f g (node_of state);
155 fun presentation_context_of state =
156 (case try node_of state of
157 SOME (Theory (_, SOME ctxt)) => ctxt
158 | SOME node => context_node node
159 | NONE => raise UNDEF);
161 fun previous_context_of (State (_, NONE)) = NONE
162 | previous_context_of (State (_, SOME prev)) = SOME (context_node prev);
164 val context_of = node_case Context.proof_of Proof.context_of;
165 val generic_theory_of = node_case I (Context.Proof o Proof.context_of);
166 val theory_of = node_case Context.theory_of Proof.theory_of;
167 val proof_of = node_case (fn _ => error "No proof state") I;
169 fun proof_position_of state =
170 (case node_of state of
171 Proof (prf, _) => Proof_Node.position prf
174 fun end_theory _ (State (NONE, SOME (Theory (Context.Theory thy, _)))) = thy
175 | end_theory pos (State (NONE, _)) = error ("Bad theory" ^ Position.here pos)
176 | end_theory pos (State (SOME _, _)) = error ("Unfinished theory" ^ Position.here pos);
181 fun pretty_context state =
182 (case try node_of state of
188 Theory (gthy, _) => gthy
189 | Proof (_, (_, gthy)) => gthy
190 | Skipped_Proof (_, (gthy, _)) => gthy);
191 val lthy = Context.cases (Named_Target.theory_init) I gthy;
192 in Local_Theory.pretty lthy end);
194 fun pretty_state state =
195 (case try node_of state of
197 | SOME (Theory _) => []
198 | SOME (Proof (prf, _)) => Proof.pretty_state (Proof_Node.current prf)
199 | SOME (Skipped_Proof (d, _)) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)]);
201 val string_of_state = pretty_state #> Pretty.chunks #> Pretty.string_of;
203 fun pretty_abstract state = Pretty.str ("<Isar " ^ str_of_state state ^ ">");
205 val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_abstract);
209 (** toplevel transitions **)
211 (* node transactions -- maintaining stable checkpoints *)
213 exception FAILURE of state * exn;
217 fun reset_presentation (Theory (gthy, _)) = Theory (gthy, NONE)
218 | reset_presentation node = node;
222 fun apply_transaction f g node =
224 val cont_node = reset_presentation node;
225 val context = cases_node I (Context.Proof o Proof.context_of) cont_node;
226 fun state_error e nd = (State (SOME nd, SOME node), e);
230 |> Runtime.controlled_execution (SOME context) f
232 handle exn => state_error (SOME exn) cont_node;
236 | SOME exn => raise FAILURE (result, exn))
239 val exit_transaction =
241 (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (Theory.end_theory thy), NONE)
242 | node => node) (K ())
243 #> (fn State (node', _) => State (NONE, node'));
248 (* primitive transitions *)
251 Init of unit -> theory | (*init theory*)
252 Exit | (*formal exit of theory*)
253 Keep of bool -> state -> unit | (*peek at state*)
254 Transaction of (bool -> node -> node) * (state -> unit); (*node transaction and presentation*)
258 fun apply_tr _ (Init f) (State (NONE, _)) =
259 State (SOME (Theory (Context.Theory (Runtime.controlled_execution NONE f ()), NONE)), NONE)
260 | apply_tr _ Exit (State (SOME (state as Theory (Context.Theory _, _)), _)) =
261 exit_transaction state
262 | apply_tr int (Keep f) state =
263 Runtime.controlled_execution (try generic_theory_of state) (fn x => tap (f int) x) state
264 | apply_tr int (Transaction (f, g)) (State (SOME state, _)) =
265 apply_transaction (fn x => f int x) g state
266 | apply_tr _ _ _ = raise UNDEF;
268 fun apply_union _ [] state = raise FAILURE (state, UNDEF)
269 | apply_union int (tr :: trs) state =
270 apply_union int trs state
271 handle Runtime.UNDEF => apply_tr int tr state
272 | FAILURE (alt_state, UNDEF) => apply_tr int tr alt_state
273 | exn as FAILURE _ => raise exn
274 | exn => raise FAILURE (state, exn);
278 fun apply_trans int trs state = (apply_union int trs state, NONE)
279 handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn);
284 (* datatype transition *)
286 datatype transition = Transition of
287 {name: string, (*command name*)
288 pos: Position.T, (*source position*)
289 timing: Time.time, (*prescient timing information*)
290 trans: trans list}; (*primitive transitions (union)*)
292 fun make_transition (name, pos, timing, trans) =
293 Transition {name = name, pos = pos, timing = timing, trans = trans};
295 fun map_transition f (Transition {name, pos, timing, trans}) =
296 make_transition (f (name, pos, timing, trans));
298 val empty = make_transition ("", Position.none, Time.zeroTime, []);
303 fun name_of (Transition {name, ...}) = name;
304 fun pos_of (Transition {pos, ...}) = pos;
306 fun command_msg msg tr =
307 msg ^ "command " ^ quote (Markup.markup Markup.keyword1 (name_of tr)) ^
308 Position.here (pos_of tr);
310 fun at_command tr = command_msg "At " tr;
311 fun type_error tr = command_msg "Bad context for " tr;
314 (* modify transitions *)
316 fun name name = map_transition (fn (_, pos, timing, trans) =>
317 (name, pos, timing, trans));
319 fun position pos = map_transition (fn (name, _, timing, trans) =>
320 (name, pos, timing, trans));
322 fun add_trans tr = map_transition (fn (name, pos, timing, trans) =>
323 (name, pos, timing, tr :: trans));
325 val reset_trans = map_transition (fn (name, pos, timing, _) =>
326 (name, pos, timing, []));
329 (* basic transitions *)
331 fun init_theory f = add_trans (Init f);
333 fun is_init (Transition {trans = [Init _], ...}) = true
336 fun modify_init f tr = if is_init tr then init_theory f (reset_trans tr) else tr;
338 val exit = add_trans Exit;
339 val keep' = add_trans o Keep;
341 fun present_transaction f g = add_trans (Transaction (f, g));
342 fun transaction f = present_transaction f (K ());
344 fun keep f = add_trans (Keep (fn _ => f));
348 if is_proof st then f st
349 else if is_skipped_proof st then ()
350 else warning "No proof state");
352 fun ignored pos = empty |> name "<ignored>" |> position pos |> keep (fn _ => ());
353 fun is_ignored tr = name_of tr = "<ignored>";
355 fun malformed pos msg =
356 empty |> name "<malformed>" |> position pos |> keep (fn _ => error msg);
359 (* theory transitions *)
361 fun generic_theory f = transaction (fn _ =>
362 (fn Theory (gthy, _) => Theory (f gthy, NONE)
363 | _ => raise UNDEF));
365 fun theory' f = transaction (fn int =>
366 (fn Theory (Context.Theory thy, _) =>
371 in Theory (Context.Theory thy', NONE) end
372 | _ => raise UNDEF));
374 fun theory f = theory' (K f);
376 fun begin_local_theory begin f = transaction (fn _ =>
377 (fn Theory (Context.Theory thy, _) =>
380 val gthy = if begin then Context.Proof lthy else Context.Theory (Named_Target.exit lthy);
382 (case Local_Theory.pretty lthy of
384 | prts => Output.state (Pretty.string_of (Pretty.chunks prts)));
385 in Theory (gthy, SOME lthy) end
386 | _ => raise UNDEF));
388 val end_local_theory = transaction (fn _ =>
389 (fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (Named_Target.exit lthy), SOME lthy)
390 | _ => raise UNDEF));
392 fun open_target f = transaction (fn _ =>
393 (fn Theory (gthy, _) =>
394 let val lthy = f gthy
395 in Theory (Context.Proof lthy, SOME lthy) end
396 | _ => raise UNDEF));
398 val close_target = transaction (fn _ =>
399 (fn Theory (Context.Proof lthy, _) =>
400 (case try Local_Theory.close_target lthy of
404 if can Local_Theory.assert ctxt'
405 then Context.Proof ctxt'
406 else Context.Theory (Proof_Context.theory_of ctxt');
407 in Theory (gthy', SOME lthy) end
408 | NONE => raise UNDEF)
409 | _ => raise UNDEF));
411 fun restricted_context (SOME (strict, scope)) =
412 Proof_Context.map_naming (Name_Space.restricted strict scope)
413 | restricted_context NONE = I;
415 fun local_theory' restricted target f = present_transaction (fn int =>
416 (fn Theory (gthy, _) =>
418 val (finish, lthy) = Named_Target.switch target gthy;
420 |> restricted_context restricted
421 |> Local_Theory.new_group
423 |> Local_Theory.reset_group;
424 in Theory (finish lthy', SOME lthy') end
428 fun local_theory restricted target f = local_theory' restricted target (K f);
430 fun present_local_theory target = present_transaction (fn _ =>
431 (fn Theory (gthy, _) =>
432 let val (finish, lthy) = Named_Target.switch target gthy;
433 in Theory (finish lthy, SOME lthy) end
434 | _ => raise UNDEF));
437 (* proof transitions *)
439 fun end_proof f = transaction (fn int =>
440 (fn Proof (prf, (finish, _)) =>
441 let val state = Proof_Node.current prf in
442 if can (Proof.assert_bottom true) state then
444 val ctxt' = f int state;
445 val gthy' = finish ctxt';
446 in Theory (gthy', SOME ctxt') end
449 | Skipped_Proof (0, (gthy, _)) => Theory (gthy, NONE)
450 | _ => raise UNDEF));
454 fun begin_proof init = transaction (fn int =>
455 (fn Theory (gthy, _) =>
457 val (finish, prf) = init int gthy;
458 val document = Options.default_string "document";
459 val skip = (document = "" orelse document = "false") andalso Goal.skip_proofs_enabled ();
460 val schematic_goal = try Proof.schematic_goal prf;
462 if skip andalso schematic_goal = SOME true then
463 warning "Cannot skip proof of schematic goal statement"
466 if skip andalso schematic_goal = SOME false then
467 Skipped_Proof (0, (finish (Proof.global_skip_proof true prf), gthy))
468 else Proof (Proof_Node.init prf, (finish, gthy))
470 | _ => raise UNDEF));
474 fun local_theory_to_proof' restricted target f = begin_proof
475 (fn int => fn gthy =>
477 val (finish, lthy) = Named_Target.switch target gthy;
479 |> restricted_context restricted
480 |> Local_Theory.new_group
482 in (finish o Local_Theory.reset_group, prf) end);
484 fun local_theory_to_proof restricted target f =
485 local_theory_to_proof' restricted target (K f);
487 fun theory_to_proof f = begin_proof
489 (Context.Theory o Sign.reset_group o Sign.change_check o Proof_Context.theory_of,
491 Context.Theory thy => f (Sign.new_group thy)
492 | _ => raise UNDEF)));
496 fun forget_proof strict = transaction (fn _ =>
497 (fn Proof (prf, (_, orig_gthy)) =>
498 if strict andalso not (Proof.has_bottom_goal (Proof_Node.current prf))
499 then raise UNDEF else Theory (orig_gthy, NONE)
500 | Skipped_Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
501 | _ => raise UNDEF));
503 fun proofs' f = transaction (fn int =>
504 (fn Proof (prf, x) => Proof (Proof_Node.applys (f int) prf, x)
505 | skip as Skipped_Proof _ => skip
506 | _ => raise UNDEF));
508 fun proof' f = proofs' ((Seq.single o Seq.Result) oo f);
509 val proofs = proofs' o K;
510 val proof = proof' o K;
515 fun actual_proof f = transaction (fn _ =>
516 (fn Proof (prf, x) => Proof (f prf, x)
517 | _ => raise UNDEF));
519 fun skip_proof f = transaction (fn _ =>
520 (fn skip as Skipped_Proof _ => (f (); skip)
521 | _ => raise UNDEF));
523 val skip_proof_open = transaction (fn _ =>
524 (fn Skipped_Proof (d, x) => Skipped_Proof (d + 1, x)
525 | _ => raise UNDEF));
527 val skip_proof_close = transaction (fn _ =>
528 (fn Skipped_Proof (0, (gthy, _)) => Theory (gthy, NONE)
529 | Skipped_Proof (d, x) => Skipped_Proof (d - 1, x)
530 | _ => raise UNDEF));
534 (** toplevel transactions **)
536 (* runtime position *)
538 fun exec_id id (tr as Transition {pos, ...}) =
539 position (Position.put_id (Document_ID.print id) pos) tr;
541 fun setmp_thread_position (Transition {pos, ...}) f x =
542 Position.setmp_thread_data pos f x;
545 (* post-transition hooks *)
549 Synchronized.var "Toplevel.hooks" ([]: (transition -> state -> state -> unit) list);
552 fun add_hook hook = Synchronized.change hooks (cons hook);
553 fun get_hooks () = Synchronized.value hooks;
558 (* apply transitions *)
560 fun get_timing (Transition {timing, ...}) = timing;
561 fun put_timing timing = map_transition (fn (name, pos, _, trans) => (name, pos, timing, trans));
565 fun app int (tr as Transition {trans, ...}) =
566 setmp_thread_position tr (fn state =>
568 val timing_start = Timing.start ();
569 val (result, opt_err) = apply_trans int trans state;
570 val timing_result = Timing.result timing_start;
572 Markup.command_timing :: (Markup.nameN, name_of tr) :: Position.properties_of (pos_of tr);
573 val _ = Timing.protocol_message timing_props timing_result;
574 in (result, Option.map (fn UNDEF => ERROR (type_error tr) | exn => exn) opt_err) end);
578 fun transition int tr st =
581 Context.setmp_generic_context (try (Context.Proof o presentation_context_of) st)
582 (fn () => app int tr st) ();
583 val opt_err' = opt_err |> Option.map
584 (fn Runtime.EXCURSION_FAIL exn_info => exn_info
585 | exn => (Runtime.exn_context (try context_of st) exn, at_command tr));
586 val _ = get_hooks () |> List.app (fn f => (try (fn () => f tr st st') (); ()));
587 in (st', opt_err') end;
592 (* managed commands *)
594 fun command_errors int tr st =
595 (case transition int tr st of
596 (st', NONE) => ([], SOME st')
597 | (_, SOME (exn, _)) => (Runtime.exn_messages exn, NONE));
599 fun command_exception int tr st =
600 (case transition int tr st of
602 | (_, SOME (exn, info)) =>
603 if Exn.is_interrupt exn then Exn.reraise exn
604 else raise Runtime.EXCURSION_FAIL (exn, info));
606 val command = command_exception false;
613 fun reset_state check trans st =
614 if check st then NONE
615 else #2 (command_errors false (trans empty) st);
619 val reset_theory = reset_state is_theory (forget_proof false);
623 (transaction (fn _ =>
624 (fn Theory (gthy, _) => Skipped_Proof (0, (gthy, gthy))
625 | _ => raise UNDEF)));
630 (* scheduled proof result *)
633 Result of transition * state |
634 Result_List of result list |
635 Result_Future of result future;
637 fun join_results (Result x) = [x]
638 | join_results (Result_List xs) = maps join_results xs
639 | join_results (Result_Future x) = join_results (Future.join x);
643 structure Result = Proof_Data
646 fun init _ = Result_List [];
649 val get_result = Result.get o Proof.context_of;
650 val put_result = Proof.map_context o Result.put;
652 fun timing_estimate elem =
653 let val trs = tl (Thy_Syntax.flat_element elem)
654 in fold (fn tr => fn t => get_timing tr + t) trs Time.zeroTime end;
656 fun proof_future_enabled estimate st =
657 (case try proof_of st of
660 not (Proof.is_relevant state) andalso
661 (if can (Proof.assert_bottom true) state
662 then Goal.future_enabled 1
663 else Goal.future_enabled 2 orelse Goal.future_enabled_timing estimate));
665 fun atom_result keywords tr st =
668 if Goal.future_enabled 1 andalso Keyword.is_diag keywords (name_of tr) then
670 {name = "Toplevel.diag", pos = pos_of tr, pri = ~1}
671 (fn () => command tr st); st)
673 in (Result (tr, st'), st') end;
677 fun element_result keywords (Thy_Syntax.Element (tr, NONE)) st = atom_result keywords tr st
678 | element_result keywords (elem as Thy_Syntax.Element (head_tr, SOME element_rest)) st =
680 val (head_result, st') = atom_result keywords head_tr st;
681 val (body_elems, end_tr) = element_rest;
682 val estimate = timing_estimate elem;
684 if not (proof_future_enabled estimate st')
687 val proof_trs = maps Thy_Syntax.flat_element body_elems @ [end_tr];
688 val (proof_results, st'') = fold_map (atom_result keywords) proof_trs st';
689 in (Result_List (head_result :: proof_results), st'') end
692 val finish = Context.Theory o Proof_Context.theory_of;
695 Proof.future_proof (fn state =>
697 {name = "Toplevel.future_proof", pos = pos_of head_tr, pri = ~1}
700 val State (SOME (Proof (prf, (_, orig_gthy))), prev) = st';
701 val prf' = Proof_Node.apply (K state) prf;
702 val (result, result_state) =
703 State (SOME (Proof (prf', (finish, orig_gthy))), prev)
704 |> fold_map (element_result keywords) body_elems ||> command end_tr;
705 in (Result_List result, presentation_context_of result_state) end))
706 #> (fn (res, state') => state' |> put_result (Result_Future res));
709 proof (future_proof #>
710 (fn state => state |> Proof.local_done_proof |> put_result (get_result state))) o
711 end_proof (fn _ => future_proof #>
712 (fn state => state |> Proof.global_done_proof |> Result.put (get_result state)));
715 |> command (head_tr |> reset_trans |> forked_proof);
716 val end_result = Result (end_tr, st'');
718 Result_List [head_result, Result.get (presentation_context_of st''), end_result];
719 in (result, st'') end