src/Pure/Isar/toplevel.ML
author wenzelm
Sun, 10 Mar 2019 00:21:34 +0100
changeset 69887 b9985133805d
parent 69886 0cb8753bdb50
child 69888 6db51f45b5f9
permissions -rw-r--r--
added semantic document markers; emulate old-style tags as "tag" markers, with subtle change of semantics for multiples tags (ever used?); tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/toplevel.ML
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
     3
26602
5534b6a6b810 made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents: 26491
diff changeset
     4
Isabelle/Isar toplevel transactions.
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
     5
*)
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
     6
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
     7
signature TOPLEVEL =
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
     8
sig
19063
049c010c8fb7 added cases_node;
wenzelm
parents: 18955
diff changeset
     9
  exception UNDEF
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    10
  type state
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
    11
  val init_toplevel: unit -> state
62895
54c2abe7e9a4 treat ROOT.ML as theory with header "theory ML_Root imports ML_Bootstrap begin";
wenzelm
parents: 62889
diff changeset
    12
  val theory_toplevel: theory -> state
7732
d62523296573 added is_toplevel;
wenzelm
parents: 7612
diff changeset
    13
  val is_toplevel: state -> bool
18589
c27c9fa9e83d hide type datatype node;
wenzelm
parents: 18564
diff changeset
    14
  val is_theory: state -> bool
c27c9fa9e83d hide type datatype node;
wenzelm
parents: 18564
diff changeset
    15
  val is_proof: state -> bool
51555
6aa64925db77 explicit Toplevel.is_skipped_proof;
wenzelm
parents: 51553
diff changeset
    16
  val is_skipped_proof: state -> bool
17076
c7effdf2e2e2 state: body context;
wenzelm
parents: 16815
diff changeset
    17
  val level: state -> int
67391
d55e52e25d9a clarified signature;
wenzelm
parents: 67390
diff changeset
    18
  val previous_theory_of: state -> theory option
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 21310
diff changeset
    19
  val context_of: state -> Proof.context
22089
d9b614dc883d added generic_theory_of;
wenzelm
parents: 22056
diff changeset
    20
  val generic_theory_of: state -> generic_theory
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    21
  val theory_of: state -> theory
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    22
  val proof_of: state -> Proof.state
18589
c27c9fa9e83d hide type datatype node;
wenzelm
parents: 18564
diff changeset
    23
  val proof_position_of: state -> int
68505
088780aa2b70 clarified default tag;
wenzelm
parents: 68130
diff changeset
    24
  val is_end_theory: state -> bool
37953
ddc3b72f9a42 simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents: 37951
diff changeset
    25
  val end_theory: Position.T -> state -> theory
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
    26
  val presentation_context: state -> Proof.context
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
    27
  val presentation_state: Proof.context -> state
56893
wenzelm
parents: 56887
diff changeset
    28
  val pretty_context: state -> Pretty.T list
56887
1ca814da47ae clarified print_state, which goes back to TTY loop before Proof General, and before separate print_context;
wenzelm
parents: 56867
diff changeset
    29
  val pretty_state: state -> Pretty.T list
61208
19118f9b939d separate panel for proof state output;
wenzelm
parents: 60895
diff changeset
    30
  val string_of_state: state -> string
37858
e1ef6b441fe7 toplevel pp for Proof.state and Toplevel.state;
wenzelm
parents: 37856
diff changeset
    31
  val pretty_abstract: state -> Pretty.T
16682
wenzelm
parents: 16658
diff changeset
    32
  type transition
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    33
  val empty: transition
27427
f6751d265cf6 added name_of;
wenzelm
parents: 26982
diff changeset
    34
  val name_of: transition -> string
28105
44054657ea56 added pos_of;
wenzelm
parents: 28103
diff changeset
    35
  val pos_of: transition -> Position.T
69877
45b2e784350a tuned signature;
wenzelm
parents: 69719
diff changeset
    36
  val timing_of: transition -> Time.time
60076
e24f59cba23c tuned messages;
wenzelm
parents: 59990
diff changeset
    37
  val type_error: transition -> string
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    38
  val name: string -> transition -> transition
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    39
  val position: Position.T -> transition -> transition
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
    40
  val markers: Input.source list -> transition -> transition
69877
45b2e784350a tuned signature;
wenzelm
parents: 69719
diff changeset
    41
  val timing: Time.time -> transition -> transition
44187
88d770052bac simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents: 44186
diff changeset
    42
  val init_theory: (unit -> theory) -> transition -> transition
88d770052bac simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents: 44186
diff changeset
    43
  val is_init: transition -> bool
44186
806f0ec1a43d simplified Toplevel.init_theory: discontinued special master argument;
wenzelm
parents: 44185
diff changeset
    44
  val modify_init: (unit -> theory) -> transition -> transition
6689
wenzelm
parents: 6664
diff changeset
    45
  val exit: transition -> transition
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    46
  val keep: (state -> unit) -> transition -> transition
7612
ba11b5db431a added keep', theory';
wenzelm
parents: 7602
diff changeset
    47
  val keep': (bool -> state -> unit) -> transition -> transition
60190
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60189
diff changeset
    48
  val keep_proof: (state -> unit) -> transition -> transition
27840
e9483ef084cc added ignored, malformed transitions;
wenzelm
parents: 27606
diff changeset
    49
  val ignored: Position.T -> transition
51268
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51241
diff changeset
    50
  val is_ignored: transition -> bool
27840
e9483ef084cc added ignored, malformed transitions;
wenzelm
parents: 27606
diff changeset
    51
  val malformed: Position.T -> string -> transition
26491
c93ff30790fe added generic_theory transaction;
wenzelm
parents: 26293
diff changeset
    52
  val generic_theory: (generic_theory -> generic_theory) -> transition -> transition
7612
ba11b5db431a added keep', theory';
wenzelm
parents: 7602
diff changeset
    53
  val theory': (bool -> theory -> theory) -> transition -> transition
49012
8686c36fa27d refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents: 49011
diff changeset
    54
  val theory: (theory -> theory) -> transition -> transition
20985
de13e2a23c8e exit_local_theory: pass interactive flag;
wenzelm
parents: 20963
diff changeset
    55
  val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
    56
  val end_local_theory: transition -> transition
47069
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
    57
  val open_target: (generic_theory -> local_theory) -> transition -> transition
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
    58
  val close_target: transition -> transition
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59923
diff changeset
    59
  val local_theory': (bool * Position.T) option -> (xstring * Position.T) option ->
59923
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
    60
    (bool -> local_theory -> local_theory) -> transition -> transition
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59923
diff changeset
    61
  val local_theory: (bool * Position.T) option -> (xstring * Position.T) option ->
59923
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
    62
    (local_theory -> local_theory) -> transition -> transition
45488
6d71d9e52369 pass positions for named targets, for formal links in the document model;
wenzelm
parents: 44304
diff changeset
    63
  val present_local_theory: (xstring * Position.T) option -> (state -> unit) ->
24453
86cf57ddf8f6 Added local_theory_to_proof'
berghofe
parents: 24295
diff changeset
    64
    transition -> transition
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59923
diff changeset
    65
  val local_theory_to_proof': (bool * Position.T) option -> (xstring * Position.T) option ->
45488
6d71d9e52369 pass positions for named targets, for formal links in the document model;
wenzelm
parents: 44304
diff changeset
    66
    (bool -> local_theory -> Proof.state) -> transition -> transition
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59923
diff changeset
    67
  val local_theory_to_proof: (bool * Position.T) option -> (xstring * Position.T) option ->
45488
6d71d9e52369 pass positions for named targets, for formal links in the document model;
wenzelm
parents: 44304
diff changeset
    68
    (local_theory -> Proof.state) -> transition -> transition
17363
046c829c075f added three_buffersN, print3;
wenzelm
parents: 17321
diff changeset
    69
  val theory_to_proof: (theory -> Proof.state) -> transition -> transition
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
    70
  val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition
68876
cefaac3d24ff no reset_proof for notepad: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68875
diff changeset
    71
  val forget_proof: transition -> transition
49863
b5fb6e7f8d81 more informative errors for 'proof' and 'apply' steps;
wenzelm
parents: 49062
diff changeset
    72
  val proofs': (bool -> Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition
17904
21c6894b5998 simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
wenzelm
parents: 17513
diff changeset
    73
  val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition
49863
b5fb6e7f8d81 more informative errors for 'proof' and 'apply' steps;
wenzelm
parents: 49062
diff changeset
    74
  val proofs: (Proof.state -> Proof.state Seq.result Seq.seq) -> transition -> transition
21177
e8228486aa03 removed checkpoint interface;
wenzelm
parents: 21007
diff changeset
    75
  val proof: (Proof.state -> Proof.state) -> transition -> transition
33390
5b499f36dd25 modernized structure Proof_Node;
wenzelm
parents: 33223
diff changeset
    76
  val actual_proof: (Proof_Node.T -> Proof_Node.T) -> transition -> transition
60693
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
    77
  val skip_proof: (unit -> unit) -> transition -> transition
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
    78
  val skip_proof_open: transition -> transition
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
    79
  val skip_proof_close: transition -> transition
52536
3a35ce87a55c tuned signature;
wenzelm
parents: 52527
diff changeset
    80
  val exec_id: Document_ID.exec -> transition -> transition
28425
25d6099179a6 export setmp_thread_position;
wenzelm
parents: 28394
diff changeset
    81
  val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
28103
b79e61861f0f simplified Toplevel.add_hook: cover successful transactions only;
wenzelm
parents: 28095
diff changeset
    82
  val add_hook: (transition -> state -> state -> unit) -> unit
59055
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
    83
  val transition: bool -> transition -> state -> state * (exn * string) option
51323
wenzelm
parents: 51322
diff changeset
    84
  val command_errors: bool -> transition -> state -> Runtime.error list * state option
51284
59a03019f3bf fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents: 51282
diff changeset
    85
  val command_exception: bool -> transition -> state -> state
56937
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
    86
  val reset_theory: state -> state option
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
    87
  val reset_proof: state -> state option
68877
33d78e5e0a00 more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68876
diff changeset
    88
  val reset_notepad: state -> state option
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
    89
  val fork_presentation: transition -> transition * transition
51332
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
    90
  type result
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
    91
  val join_results: result -> (transition * state) list
68839
d8251a61cce8 clarified modules;
wenzelm
parents: 68772
diff changeset
    92
  val element_result: Keyword.keywords -> transition Thy_Element.element -> state -> result * state
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    93
end;
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    94
6965
a766de752996 fixed interrupts (eliminated races);
wenzelm
parents: 6689
diff changeset
    95
structure Toplevel: TOPLEVEL =
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    96
struct
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    97
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    98
(** toplevel state **)
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
    99
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents: 31431
diff changeset
   100
exception UNDEF = Runtime.UNDEF;
19063
049c010c8fb7 added cases_node;
wenzelm
parents: 18955
diff changeset
   101
049c010c8fb7 added cases_node;
wenzelm
parents: 18955
diff changeset
   102
21958
9dfd1ca4c0a0 refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents: 21861
diff changeset
   103
(* datatype node *)
21294
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   104
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   105
datatype node =
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   106
  Toplevel
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   107
    (*toplevel outside of theory body*) |
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   108
  Theory of generic_theory
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   109
    (*global or local theory*) |
33390
5b499f36dd25 modernized structure Proof_Node;
wenzelm
parents: 33223
diff changeset
   110
  Proof of Proof_Node.T * ((Proof.context -> generic_theory) * generic_theory)
27576
7afff36043e6 eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents: 27564
diff changeset
   111
    (*proof node, finish, original theory*) |
51555
6aa64925db77 explicit Toplevel.is_skipped_proof;
wenzelm
parents: 51553
diff changeset
   112
  Skipped_Proof of int * (generic_theory * generic_theory);
27564
fc6d34e49e17 replaced obsolete ProofHistory by ProofNode (backtracking only);
wenzelm
parents: 27500
diff changeset
   113
    (*proof depth, resulting theory, original theory*)
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   114
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   115
val theory_node = fn Theory gthy => SOME gthy | _ => NONE;
18589
c27c9fa9e83d hide type datatype node;
wenzelm
parents: 18564
diff changeset
   116
val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE;
51555
6aa64925db77 explicit Toplevel.is_skipped_proof;
wenzelm
parents: 51553
diff changeset
   117
val skipped_proof_node = fn Skipped_Proof _ => true | _ => false;
18589
c27c9fa9e83d hide type datatype node;
wenzelm
parents: 18564
diff changeset
   118
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   119
fun cases_node f _ _ Toplevel = f ()
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   120
  | cases_node _ g _ (Theory gthy) = g gthy
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   121
  | cases_node _ _ h (Proof (prf, _)) = h (Proof_Node.current prf)
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   122
  | cases_node _ g _ (Skipped_Proof (_, (gthy, _))) = g gthy;
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   123
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   124
fun cases_proper_node g h = cases_node (fn () => raise UNDEF) g h;
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   125
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   126
val get_theory = cases_node (K NONE) (SOME o Context.theory_of) (SOME o Proof.theory_of);
19063
049c010c8fb7 added cases_node;
wenzelm
parents: 18955
diff changeset
   127
21958
9dfd1ca4c0a0 refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents: 21861
diff changeset
   128
9dfd1ca4c0a0 refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents: 21861
diff changeset
   129
(* datatype state *)
9dfd1ca4c0a0 refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents: 21861
diff changeset
   130
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   131
type node_presentation = node * Proof.context;
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   132
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   133
fun init_presentation () =
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   134
  Proof_Context.init_global (Theory.get_pure_bootstrap ());
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   135
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   136
fun node_presentation node =
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   137
  (node, cases_node init_presentation Context.proof_of Proof.context_of node);
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   138
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   139
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   140
datatype state =
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   141
  State of node_presentation * theory option;
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   142
    (*current node with presentation context, previous theory*)
62895
54c2abe7e9a4 treat ROOT.ML as theory with header "theory ML_Root imports ML_Bootstrap begin";
wenzelm
parents: 62889
diff changeset
   143
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   144
fun node_of (State ((node, _), _)) = node;
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   145
fun previous_theory_of (State (_, prev_thy)) = prev_thy;
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   146
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   147
fun init_toplevel () = State (node_presentation Toplevel, NONE);
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   148
fun theory_toplevel thy = State (node_presentation (Theory (Context.Theory thy)), NONE);
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   149
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   150
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   151
fun level state =
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   152
  (case node_of state of
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   153
    Toplevel => 0
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   154
  | Theory _ => 0
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   155
  | Proof (prf, _) => Proof.level (Proof_Node.current prf)
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   156
  | Skipped_Proof (d, _) => d + 1);   (*different notion of proof depth!*)
17076
c7effdf2e2e2 state: body context;
wenzelm
parents: 16815
diff changeset
   157
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   158
fun str_of_state state =
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   159
  (case node_of state of
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   160
    Toplevel =>
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   161
      (case previous_theory_of state of
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   162
        NONE => "at top level"
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   163
      | SOME thy => "at top level, result theory " ^ quote (Context.theory_name thy))
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   164
  | Theory (Context.Theory _) => "in theory mode"
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   165
  | Theory (Context.Proof _) => "in local theory mode"
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   166
  | Proof _ => "in proof mode"
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   167
  | Skipped_Proof _ => "in skipped proof mode");
5946
a4600d21b59b print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents: 5939
diff changeset
   168
a4600d21b59b print_state hook, obeys Goals.current_goals_markers by default;
wenzelm
parents: 5939
diff changeset
   169
27576
7afff36043e6 eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents: 27564
diff changeset
   170
(* current node *)
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   171
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   172
fun is_toplevel state = (case node_of state of Toplevel => true | _ => false);
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   173
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   174
fun is_theory state =
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   175
  not (is_toplevel state) andalso is_some (theory_node (node_of state));
18589
c27c9fa9e83d hide type datatype node;
wenzelm
parents: 18564
diff changeset
   176
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   177
fun is_proof state =
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   178
  not (is_toplevel state) andalso is_some (proof_node (node_of state));
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   179
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   180
fun is_skipped_proof state =
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   181
  not (is_toplevel state) andalso skipped_proof_node (node_of state);
30801
9bdf001bea58 added Toplevel.previous_node_of;
wenzelm
parents: 30618
diff changeset
   182
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   183
fun proper_node_of state = if is_toplevel state then raise UNDEF else node_of state;
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   184
fun proper_node_case f g state = cases_proper_node f g (proper_node_of state);
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   185
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   186
val context_of = proper_node_case Context.proof_of Proof.context_of;
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   187
val generic_theory_of = proper_node_case I (Context.Proof o Proof.context_of);
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   188
val theory_of = proper_node_case Context.theory_of Proof.theory_of;
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   189
val proof_of = proper_node_case (fn _ => error "No proof state") I;
17208
49bc1bdc7b6e added no_body_context;
wenzelm
parents: 17115
diff changeset
   190
18589
c27c9fa9e83d hide type datatype node;
wenzelm
parents: 18564
diff changeset
   191
fun proof_position_of state =
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   192
  (case proper_node_of state of
33390
5b499f36dd25 modernized structure Proof_Node;
wenzelm
parents: 33223
diff changeset
   193
    Proof (prf, _) => Proof_Node.position prf
60094
96a4765ba7d1 explicit error for Toplevel.proof_of;
wenzelm
parents: 60076
diff changeset
   194
  | _ => ~1);
6664
f679ddd1ddd8 cleaned comments;
wenzelm
parents: 6244
diff changeset
   195
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   196
fun is_end_theory (State ((Toplevel, _), SOME _)) = true
68505
088780aa2b70 clarified default tag;
wenzelm
parents: 68130
diff changeset
   197
  | is_end_theory _ = false;
088780aa2b70 clarified default tag;
wenzelm
parents: 68130
diff changeset
   198
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   199
fun end_theory _ (State ((Toplevel, _), SOME thy)) = thy
68869
3739acbc2178 clarified message;
wenzelm
parents: 68839
diff changeset
   200
  | end_theory pos _ = error ("Malformed theory" ^ Position.here pos);
37953
ddc3b72f9a42 simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents: 37951
diff changeset
   201
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   202
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   203
(* presentation context *)
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   204
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   205
structure Presentation_State = Proof_Data
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   206
(
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   207
  type T = state option;
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   208
  fun init _ = NONE;
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   209
);
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   210
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   211
fun presentation_context0 (State ((_, pr_ctxt), _)) = pr_ctxt;
67642
10ff1f077119 more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents: 67641
diff changeset
   212
10ff1f077119 more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents: 67641
diff changeset
   213
fun presentation_context (state as State (current, _)) =
10ff1f077119 more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents: 67641
diff changeset
   214
  presentation_context0 state
10ff1f077119 more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents: 67641
diff changeset
   215
  |> Presentation_State.put (SOME (State (current, NONE)));
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   216
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   217
fun presentation_state ctxt =
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   218
  (case Presentation_State.get ctxt of
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   219
    NONE => State (node_presentation (Theory (Context.Proof ctxt)), NONE)
67381
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   220
  | SOME state => state);
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   221
146757999c8d theory Pure is default presentation context;
wenzelm
parents: 67376
diff changeset
   222
16815
13d20ed9086c added print_state_hook;
wenzelm
parents: 16729
diff changeset
   223
(* print state *)
13d20ed9086c added print_state_hook;
wenzelm
parents: 16729
diff changeset
   224
56893
wenzelm
parents: 56887
diff changeset
   225
fun pretty_context state =
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   226
  if is_toplevel state then []
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   227
  else
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   228
    let
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   229
      val gthy =
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   230
        (case node_of state of
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   231
          Toplevel => raise Match
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   232
        | Theory gthy => gthy
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   233
        | Proof (_, (_, gthy)) => gthy
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   234
        | Skipped_Proof (_, (_, gthy)) => gthy);
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   235
      val lthy = Context.cases Named_Target.theory_init I gthy;
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   236
    in Local_Theory.pretty lthy end;
16815
13d20ed9086c added print_state_hook;
wenzelm
parents: 16729
diff changeset
   237
56887
1ca814da47ae clarified print_state, which goes back to TTY loop before Proof General, and before separate print_context;
wenzelm
parents: 56867
diff changeset
   238
fun pretty_state state =
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   239
  (case node_of state of
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   240
    Toplevel => []
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   241
  | Theory _ => []
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   242
  | Proof (prf, _) => Proof.pretty_state (Proof_Node.current prf)
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   243
  | Skipped_Proof (d, _) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)]);
56867
224109105008 more print operations;
wenzelm
parents: 56334
diff changeset
   244
61208
19118f9b939d separate panel for proof state output;
wenzelm
parents: 60895
diff changeset
   245
val string_of_state = pretty_state #> Pretty.chunks #> Pretty.string_of;
16815
13d20ed9086c added print_state_hook;
wenzelm
parents: 16729
diff changeset
   246
37858
e1ef6b441fe7 toplevel pp for Proof.state and Toplevel.state;
wenzelm
parents: 37856
diff changeset
   247
fun pretty_abstract state = Pretty.str ("<Isar " ^ str_of_state state ^ ">");
e1ef6b441fe7 toplevel pp for Proof.state and Toplevel.state;
wenzelm
parents: 37856
diff changeset
   248
62819
d3ff367a16a0 careful export of type-dependent functions, without losing their special status;
wenzelm
parents: 62663
diff changeset
   249
val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_abstract);
62663
bea354f6ff21 clarified modules;
wenzelm
parents: 62505
diff changeset
   250
16815
13d20ed9086c added print_state_hook;
wenzelm
parents: 16729
diff changeset
   251
15668
53c049a365cf improved exn_message;
wenzelm
parents: 15633
diff changeset
   252
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   253
(** toplevel transitions **)
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   254
27601
6683cdb94af8 simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents: 27583
diff changeset
   255
(* node transactions -- maintaining stable checkpoints *)
7022
abf9d5e2fb6e removed BREAK, ROLLBACK;
wenzelm
parents: 6971
diff changeset
   256
31476
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents: 31431
diff changeset
   257
exception FAILURE of state * exn;
c5d2899b6de9 moved Isar toplevel runtime support to runtime.ML, which is loaded early (before ml_compiler.ML);
wenzelm
parents: 31431
diff changeset
   258
37953
ddc3b72f9a42 simplified handling of theory begin/end wrt. toplevel and theory loader;
wenzelm
parents: 37951
diff changeset
   259
fun apply_transaction f g node =
20128
8f0e07d7cf92 keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents: 19996
diff changeset
   260
  let
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   261
    val node_pr = node_presentation node;
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   262
    val context = cases_proper_node I (Context.Proof o Proof.context_of) node;
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   263
    fun state_error e node_pr' = (State (node_pr', get_theory node), e);
6689
wenzelm
parents: 6664
diff changeset
   264
20128
8f0e07d7cf92 keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents: 19996
diff changeset
   265
    val (result, err) =
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   266
      node
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 56265
diff changeset
   267
      |> Runtime.controlled_execution (SOME context) f
26624
770265032999 transaction/init: ensure stable theory (non-draft);
wenzelm
parents: 26621
diff changeset
   268
      |> state_error NONE
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   269
      handle exn => state_error (SOME exn) node_pr;
20128
8f0e07d7cf92 keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents: 19996
diff changeset
   270
  in
52696
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52565
diff changeset
   271
    (case err of
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52565
diff changeset
   272
      NONE => tap g result
38466f4f3483 immutable theory values with full stamp record of every update (increase of stamp size for HOL: 20000 -> 100000, JinjaThreads: 65000 -> 300000) -- minimal measurable impact on inference kernel performance;
wenzelm
parents: 52565
diff changeset
   273
    | SOME exn => raise FAILURE (result, exn))
20128
8f0e07d7cf92 keep/transaction: unified execution model (with debugging etc.);
wenzelm
parents: 19996
diff changeset
   274
  end;
6689
wenzelm
parents: 6664
diff changeset
   275
wenzelm
parents: 6664
diff changeset
   276
wenzelm
parents: 6664
diff changeset
   277
(* primitive transitions *)
wenzelm
parents: 6664
diff changeset
   278
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   279
datatype trans =
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   280
  (*init theory*)
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   281
  Init of unit -> theory |
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   282
  (*formal exit of theory*)
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   283
  Exit |
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   284
  (*peek at state*)
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   285
  Keep of bool -> state -> unit |
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   286
  (*node transaction and presentation*)
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   287
  Transaction of (bool -> node -> node_presentation) * (state -> unit);
21958
9dfd1ca4c0a0 refined notion of empty toplevel, admits undo of 'end';
wenzelm
parents: 21861
diff changeset
   288
6689
wenzelm
parents: 6664
diff changeset
   289
local
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   290
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   291
fun apply_tr int trans state =
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   292
  (case (trans, node_of state) of
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   293
    (Init f, Toplevel) =>
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   294
      Runtime.controlled_execution NONE (fn () =>
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   295
        State (node_presentation (Theory (Context.Theory (f ()))), NONE)) ()
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   296
  | (Exit, node as Theory (Context.Theory thy)) =>
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   297
      let
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   298
        val State ((node', pr_ctxt), _) =
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   299
          node |> apply_transaction
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   300
            (fn _ => node_presentation (Theory (Context.Theory (Theory.end_theory thy))))
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   301
            (K ());
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   302
      in State ((Toplevel, pr_ctxt), get_theory node') end
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   303
  | (Keep f, node) =>
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   304
      Runtime.controlled_execution (try generic_theory_of state)
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   305
        (fn () => (f int state; State (node_presentation node, previous_theory_of state))) ()
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   306
  | (Transaction _, Toplevel) => raise UNDEF
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   307
  | (Transaction (f, g), node) => apply_transaction (fn x => f int x) g node
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   308
  | _ => raise UNDEF);
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   309
32792
a08a2b962a09 eliminated dead code;
wenzelm
parents: 32738
diff changeset
   310
fun apply_union _ [] state = raise FAILURE (state, UNDEF)
a08a2b962a09 eliminated dead code;
wenzelm
parents: 32738
diff changeset
   311
  | apply_union int (tr :: trs) state =
a08a2b962a09 eliminated dead code;
wenzelm
parents: 32738
diff changeset
   312
      apply_union int trs state
a08a2b962a09 eliminated dead code;
wenzelm
parents: 32738
diff changeset
   313
        handle Runtime.UNDEF => apply_tr int tr state
a08a2b962a09 eliminated dead code;
wenzelm
parents: 32738
diff changeset
   314
          | FAILURE (alt_state, UNDEF) => apply_tr int tr alt_state
6689
wenzelm
parents: 6664
diff changeset
   315
          | exn as FAILURE _ => raise exn
wenzelm
parents: 6664
diff changeset
   316
          | exn => raise FAILURE (state, exn);
wenzelm
parents: 6664
diff changeset
   317
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   318
fun apply_markers markers (state as State ((node, pr_ctxt), prev_thy)) =
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   319
  let
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   320
    val state' =
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   321
      Runtime.controlled_execution (try generic_theory_of state)
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   322
        (fn () => State ((node, fold Document_Marker.evaluate markers pr_ctxt), prev_thy)) ();
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   323
  in (state', NONE) end
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   324
  handle exn => (state, SOME exn);
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   325
6689
wenzelm
parents: 6664
diff changeset
   326
in
wenzelm
parents: 6664
diff changeset
   327
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   328
fun apply_trans int trans markers state =
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   329
  (apply_union int trans state |> apply_markers markers)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15519
diff changeset
   330
  handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn);
6689
wenzelm
parents: 6664
diff changeset
   331
wenzelm
parents: 6664
diff changeset
   332
end;
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   333
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   334
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   335
(* datatype transition *)
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   336
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   337
datatype transition = Transition of
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   338
 {name: string,               (*command name*)
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   339
  pos: Position.T,            (*source position*)
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   340
  markers: Input.source list, (*semantic document markers*)
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   341
  timing: Time.time,          (*prescient timing information*)
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   342
  trans: trans list};         (*primitive transitions (union)*)
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   343
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   344
fun make_transition (name, pos, markers, timing, trans) =
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   345
  Transition {name = name, pos = pos, markers = markers, timing = timing, trans = trans};
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   346
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   347
fun map_transition f (Transition {name, pos, markers, timing, trans}) =
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   348
  make_transition (f (name, pos, markers, timing, trans));
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   349
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   350
val empty = make_transition ("", Position.none, [], Time.zeroTime, []);
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   351
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   352
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   353
(* diagnostics *)
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   354
27427
f6751d265cf6 added name_of;
wenzelm
parents: 26982
diff changeset
   355
fun name_of (Transition {name, ...}) = name;
28105
44054657ea56 added pos_of;
wenzelm
parents: 28103
diff changeset
   356
fun pos_of (Transition {pos, ...}) = pos;
69877
45b2e784350a tuned signature;
wenzelm
parents: 69719
diff changeset
   357
fun timing_of (Transition {timing, ...}) = timing;
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   358
60076
e24f59cba23c tuned messages;
wenzelm
parents: 59990
diff changeset
   359
fun command_msg msg tr =
e24f59cba23c tuned messages;
wenzelm
parents: 59990
diff changeset
   360
  msg ^ "command " ^ quote (Markup.markup Markup.keyword1 (name_of tr)) ^
e24f59cba23c tuned messages;
wenzelm
parents: 59990
diff changeset
   361
    Position.here (pos_of tr);
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   362
60076
e24f59cba23c tuned messages;
wenzelm
parents: 59990
diff changeset
   363
fun at_command tr = command_msg "At " tr;
e24f59cba23c tuned messages;
wenzelm
parents: 59990
diff changeset
   364
fun type_error tr = command_msg "Bad context for " tr;
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   365
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   366
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   367
(* modify transitions *)
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   368
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   369
fun name name = map_transition (fn (_, pos, markers, timing, trans) =>
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   370
  (name, pos, markers, timing, trans));
9010
ce78dc5e1a73 Toplevel.no_timing;
wenzelm
parents: 8999
diff changeset
   371
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   372
fun position pos = map_transition (fn (name, _, markers, timing, trans) =>
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   373
  (name, pos, markers, timing, trans));
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   374
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   375
fun markers markers = map_transition (fn (name, pos, _, timing, trans) =>
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   376
  (name, pos, markers, timing, trans));
14923
8a32071c3c13 added name_of, source_of, source;
wenzelm
parents: 14825
diff changeset
   377
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   378
fun timing timing = map_transition (fn (name, pos, markers, _, trans) =>
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   379
  (name, pos, markers, timing, trans));
69877
45b2e784350a tuned signature;
wenzelm
parents: 69719
diff changeset
   380
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   381
fun add_trans tr = map_transition (fn (name, pos, markers, timing, trans) =>
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   382
  (name, pos, markers, timing, tr :: trans));
16607
81e687c63e29 added print': print depending on print_mode;
wenzelm
parents: 16490
diff changeset
   383
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   384
val reset_trans = map_transition (fn (name, pos, markers, timing, _) =>
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   385
  (name, pos, markers, timing, []));
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   386
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   387
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   388
(* basic transitions *)
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   389
44187
88d770052bac simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents: 44186
diff changeset
   390
fun init_theory f = add_trans (Init f);
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37953
diff changeset
   391
44187
88d770052bac simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents: 44186
diff changeset
   392
fun is_init (Transition {trans = [Init _], ...}) = true
88d770052bac simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents: 44186
diff changeset
   393
  | is_init _ = false;
88d770052bac simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents: 44186
diff changeset
   394
88d770052bac simplified Toplevel.init_theory: discontinued special name argument;
wenzelm
parents: 44186
diff changeset
   395
fun modify_init f tr = if is_init tr then init_theory f (reset_trans tr) else tr;
37977
3ceccd415145 simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
wenzelm
parents: 37953
diff changeset
   396
6689
wenzelm
parents: 6664
diff changeset
   397
val exit = add_trans Exit;
7612
ba11b5db431a added keep', theory';
wenzelm
parents: 7602
diff changeset
   398
val keep' = add_trans o Keep;
30366
e3d788b9dffb simplified presentation: built into transaction, pass state directly;
wenzelm
parents: 29516
diff changeset
   399
e3d788b9dffb simplified presentation: built into transaction, pass state directly;
wenzelm
parents: 29516
diff changeset
   400
fun present_transaction f g = add_trans (Transaction (f, g));
e3d788b9dffb simplified presentation: built into transaction, pass state directly;
wenzelm
parents: 29516
diff changeset
   401
fun transaction f = present_transaction f (K ());
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   402
fun transaction0 f = present_transaction (node_presentation oo f) (K ());
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   403
7612
ba11b5db431a added keep', theory';
wenzelm
parents: 7602
diff changeset
   404
fun keep f = add_trans (Keep (fn _ => f));
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   405
60190
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60189
diff changeset
   406
fun keep_proof f =
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60189
diff changeset
   407
  keep (fn st =>
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60189
diff changeset
   408
    if is_proof st then f st
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60189
diff changeset
   409
    else if is_skipped_proof st then ()
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60189
diff changeset
   410
    else warning "No proof state");
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60189
diff changeset
   411
60189
0d3a62127057 tuned signature;
wenzelm
parents: 60097
diff changeset
   412
fun ignored pos = empty |> name "<ignored>" |> position pos |> keep (fn _ => ());
51268
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51241
diff changeset
   413
fun is_ignored tr = name_of tr = "<ignored>";
48772
e46cd0d26481 vacuous execution after first malformed command;
wenzelm
parents: 48633
diff changeset
   414
27840
e9483ef084cc added ignored, malformed transitions;
wenzelm
parents: 27606
diff changeset
   415
fun malformed pos msg =
60189
0d3a62127057 tuned signature;
wenzelm
parents: 60097
diff changeset
   416
  empty |> name "<malformed>" |> position pos |> keep (fn _ => error msg);
27840
e9483ef084cc added ignored, malformed transitions;
wenzelm
parents: 27606
diff changeset
   417
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   418
49012
8686c36fa27d refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents: 49011
diff changeset
   419
(* theory transitions *)
44304
7ee000ce5390 maintain recent future proofs at transaction boundaries;
wenzelm
parents: 44270
diff changeset
   420
27601
6683cdb94af8 simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents: 27583
diff changeset
   421
fun generic_theory f = transaction (fn _ =>
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   422
  (fn Theory gthy => node_presentation (Theory (f gthy))
26491
c93ff30790fe added generic_theory transaction;
wenzelm
parents: 26293
diff changeset
   423
    | _ => raise UNDEF));
c93ff30790fe added generic_theory transaction;
wenzelm
parents: 26293
diff changeset
   424
27601
6683cdb94af8 simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents: 27583
diff changeset
   425
fun theory' f = transaction (fn int =>
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   426
  (fn Theory (Context.Theory thy) =>
33725
a8481da77270 implicit name space grouping for theory/local_theory transactions;
wenzelm
parents: 33671
diff changeset
   427
      let val thy' = thy
49012
8686c36fa27d refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents: 49011
diff changeset
   428
        |> Sign.new_group
33725
a8481da77270 implicit name space grouping for theory/local_theory transactions;
wenzelm
parents: 33671
diff changeset
   429
        |> f int
a8481da77270 implicit name space grouping for theory/local_theory transactions;
wenzelm
parents: 33671
diff changeset
   430
        |> Sign.reset_group;
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   431
      in node_presentation (Theory (Context.Theory thy')) end
20963
a7fd8f05a2be added type global_theory -- theory or local_theory;
wenzelm
parents: 20928
diff changeset
   432
    | _ => raise UNDEF));
a7fd8f05a2be added type global_theory -- theory or local_theory;
wenzelm
parents: 20928
diff changeset
   433
a7fd8f05a2be added type global_theory -- theory or local_theory;
wenzelm
parents: 20928
diff changeset
   434
fun theory f = theory' (K f);
a7fd8f05a2be added type global_theory -- theory or local_theory;
wenzelm
parents: 20928
diff changeset
   435
27601
6683cdb94af8 simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents: 27583
diff changeset
   436
fun begin_local_theory begin f = transaction (fn _ =>
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   437
  (fn Theory (Context.Theory thy) =>
20963
a7fd8f05a2be added type global_theory -- theory or local_theory;
wenzelm
parents: 20928
diff changeset
   438
        let
20985
de13e2a23c8e exit_local_theory: pass interactive flag;
wenzelm
parents: 20963
diff changeset
   439
          val lthy = f thy;
57483
950dc7446454 centralized (ad-hoc) switching of targets in named_target.ML
haftmann
parents: 57184
diff changeset
   440
          val gthy = if begin then Context.Proof lthy else Context.Theory (Named_Target.exit lthy);
56897
c668735fb8b5 print results as "state", to avoid intrusion into the source text;
wenzelm
parents: 56895
diff changeset
   441
          val _ =
60245
79ad597fe699 tuned output;
wenzelm
parents: 60190
diff changeset
   442
            (case Local_Theory.pretty lthy of
79ad597fe699 tuned output;
wenzelm
parents: 60190
diff changeset
   443
              [] => ()
79ad597fe699 tuned output;
wenzelm
parents: 60190
diff changeset
   444
            | prts => Output.state (Pretty.string_of (Pretty.chunks prts)));
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   445
        in (Theory gthy, lthy) end
20963
a7fd8f05a2be added type global_theory -- theory or local_theory;
wenzelm
parents: 20928
diff changeset
   446
    | _ => raise UNDEF));
17076
c7effdf2e2e2 state: body context;
wenzelm
parents: 16815
diff changeset
   447
27601
6683cdb94af8 simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents: 27583
diff changeset
   448
val end_local_theory = transaction (fn _ =>
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   449
  (fn Theory (Context.Proof lthy) => (Theory (Context.Theory (Named_Target.exit lthy)), lthy)
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   450
    | _ => raise UNDEF));
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   451
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   452
fun open_target f = transaction0 (fn _ =>
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   453
  (fn Theory gthy =>
47069
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
   454
        let val lthy = f gthy
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   455
        in Theory (Context.Proof lthy) end
47069
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
   456
    | _ => raise UNDEF));
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
   457
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
   458
val close_target = transaction (fn _ =>
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   459
  (fn Theory (Context.Proof lthy) =>
47069
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
   460
        (case try Local_Theory.close_target lthy of
50739
5165d7e6d3b9 more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents: 50201
diff changeset
   461
          SOME ctxt' =>
5165d7e6d3b9 more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents: 50201
diff changeset
   462
            let
5165d7e6d3b9 more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents: 50201
diff changeset
   463
              val gthy' =
5165d7e6d3b9 more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents: 50201
diff changeset
   464
                if can Local_Theory.assert ctxt'
5165d7e6d3b9 more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents: 50201
diff changeset
   465
                then Context.Proof ctxt'
5165d7e6d3b9 more precise Local_Theory.level: 1 really means main target and >= 2 nested context;
wenzelm
parents: 50201
diff changeset
   466
                else Context.Theory (Proof_Context.theory_of ctxt');
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   467
            in (Theory gthy', lthy) end
47069
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
   468
        | NONE => raise UNDEF)
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
   469
    | _ => raise UNDEF));
451fc10a81f3 more explicit Toplevel.open_target/close_target;
wenzelm
parents: 46959
diff changeset
   470
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   471
fun restricted_context (SOME (strict, scope)) =
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   472
      Proof_Context.map_naming (Name_Space.restricted strict scope)
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   473
  | restricted_context NONE = I;
59939
7d46aa03696e support for 'restricted' modifier: only qualified accesses outside the local scope;
wenzelm
parents: 59923
diff changeset
   474
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   475
fun local_theory' restricted target f = present_transaction (fn int =>
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   476
  (fn Theory gthy =>
21294
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   477
        let
59923
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
   478
          val (finish, lthy) = Named_Target.switch target gthy;
47274
feef9b0b6031 misc tuning and simplification;
wenzelm
parents: 47069
diff changeset
   479
          val lthy' = lthy
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   480
            |> restricted_context restricted
49012
8686c36fa27d refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents: 49011
diff changeset
   481
            |> Local_Theory.new_group
33725
a8481da77270 implicit name space grouping for theory/local_theory transactions;
wenzelm
parents: 33671
diff changeset
   482
            |> f int
a8481da77270 implicit name space grouping for theory/local_theory transactions;
wenzelm
parents: 33671
diff changeset
   483
            |> Local_Theory.reset_group;
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   484
        in (Theory (finish lthy'), lthy') end
59032
wenzelm
parents: 58999
diff changeset
   485
    | _ => raise UNDEF))
wenzelm
parents: 58999
diff changeset
   486
  (K ());
15668
53c049a365cf improved exn_message;
wenzelm
parents: 15633
diff changeset
   487
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   488
fun local_theory restricted target f = local_theory' restricted target (K f);
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   489
65054
wenzelm
parents: 62895
diff changeset
   490
fun present_local_theory target = present_transaction (fn _ =>
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   491
  (fn Theory gthy =>
59923
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
   492
        let val (finish, lthy) = Named_Target.switch target gthy;
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   493
        in (Theory (finish lthy), lthy) end
59032
wenzelm
parents: 58999
diff changeset
   494
    | _ => raise UNDEF));
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   495
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   496
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   497
(* proof transitions *)
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   498
27601
6683cdb94af8 simplified commit_exit: operate on previous node of final state, include warning here;
wenzelm
parents: 27583
diff changeset
   499
fun end_proof f = transaction (fn int =>
24795
6f5cb7885fd7 print_state_context: local theory context, not proof context;
wenzelm
parents: 24780
diff changeset
   500
  (fn Proof (prf, (finish, _)) =>
33390
5b499f36dd25 modernized structure Proof_Node;
wenzelm
parents: 33223
diff changeset
   501
        let val state = Proof_Node.current prf in
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   502
          if can (Proof.assert_bottom true) state then
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   503
            let
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   504
              val ctxt' = f int state;
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   505
              val gthy' = finish ctxt';
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   506
            in (Theory gthy', ctxt') end
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   507
          else raise UNDEF
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   508
        end
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   509
    | Skipped_Proof (0, (gthy, _)) => node_presentation (Theory gthy)
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   510
    | _ => raise UNDEF));
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   511
21294
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   512
local
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   513
69882
wenzelm
parents: 69878
diff changeset
   514
fun begin_proof init_proof = transaction0 (fn int =>
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   515
  (fn Theory gthy =>
21294
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   516
    let
69882
wenzelm
parents: 69878
diff changeset
   517
      val (finish, prf) = init_proof int gthy;
67157
d0657c8b7616 clarified document preparation vs. skip_proofs;
wenzelm
parents: 66169
diff changeset
   518
      val document = Options.default_string "document";
d0657c8b7616 clarified document preparation vs. skip_proofs;
wenzelm
parents: 66169
diff changeset
   519
      val skip = (document = "" orelse document = "false") andalso Goal.skip_proofs_enabled ();
58795
wenzelm
parents: 57483
diff changeset
   520
      val schematic_goal = try Proof.schematic_goal prf;
47274
feef9b0b6031 misc tuning and simplification;
wenzelm
parents: 47069
diff changeset
   521
      val _ =
58795
wenzelm
parents: 57483
diff changeset
   522
        if skip andalso schematic_goal = SOME true then
47274
feef9b0b6031 misc tuning and simplification;
wenzelm
parents: 47069
diff changeset
   523
          warning "Cannot skip proof of schematic goal statement"
feef9b0b6031 misc tuning and simplification;
wenzelm
parents: 47069
diff changeset
   524
        else ();
21294
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   525
    in
58795
wenzelm
parents: 57483
diff changeset
   526
      if skip andalso schematic_goal = SOME false then
51555
6aa64925db77 explicit Toplevel.is_skipped_proof;
wenzelm
parents: 51553
diff changeset
   527
        Skipped_Proof (0, (finish (Proof.global_skip_proof true prf), gthy))
47274
feef9b0b6031 misc tuning and simplification;
wenzelm
parents: 47069
diff changeset
   528
      else Proof (Proof_Node.init prf, (finish, gthy))
21294
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   529
    end
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   530
  | _ => raise UNDEF));
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   531
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   532
in
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   533
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   534
fun local_theory_to_proof' restricted target f = begin_proof
47274
feef9b0b6031 misc tuning and simplification;
wenzelm
parents: 47069
diff changeset
   535
  (fn int => fn gthy =>
59923
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
   536
    let
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
   537
      val (finish, lthy) = Named_Target.switch target gthy;
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
   538
      val prf = lthy
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   539
        |> restricted_context restricted
59923
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
   540
        |> Local_Theory.new_group
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
   541
        |> f int;
b21c82422d65 support private scope for individual local theory commands;
wenzelm
parents: 59472
diff changeset
   542
    in (finish o Local_Theory.reset_group, prf) end);
24780
47bb1e380d83 local_theory transactions: more careful treatment of context position;
wenzelm
parents: 24634
diff changeset
   543
59990
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   544
fun local_theory_to_proof restricted target f =
a81dc82ecba3 clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents: 59939
diff changeset
   545
  local_theory_to_proof' restricted target (K f);
21294
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   546
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   547
fun theory_to_proof f = begin_proof
47274
feef9b0b6031 misc tuning and simplification;
wenzelm
parents: 47069
diff changeset
   548
  (fn _ => fn gthy =>
56057
ad6bd8030d88 more explicit Sign.change_check -- detect structural mistakes where they emerge, not at later theory merges;
wenzelm
parents: 56055
diff changeset
   549
    (Context.Theory o Sign.reset_group o Sign.change_check o Proof_Context.theory_of,
49062
7e31dfd99ce7 discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents: 49042
diff changeset
   550
      (case gthy of
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52696
diff changeset
   551
        Context.Theory thy => f (Sign.new_group thy)
49012
8686c36fa27d refined treatment of forked proofs at transaction boundaries, including proof commands (see also 7ee000ce5390);
wenzelm
parents: 49011
diff changeset
   552
      | _ => raise UNDEF)));
21294
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   553
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   554
end;
5cd48242ef17 simplified local theory wrappers;
wenzelm
parents: 21277
diff changeset
   555
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   556
val forget_proof = transaction0 (fn _ =>
58798
49ed5eea15d4 'oops' requires proper goal statement -- exclude 'notepad' to avoid disrupting begin/end structure;
wenzelm
parents: 58795
diff changeset
   557
  (fn Proof (prf, (_, orig_gthy)) =>
68876
cefaac3d24ff no reset_proof for notepad: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68875
diff changeset
   558
        if Proof.is_notepad (Proof_Node.current prf) then raise UNDEF
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   559
        else Theory orig_gthy
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   560
    | Skipped_Proof (_, (_, orig_gthy)) => Theory orig_gthy
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   561
    | _ => raise UNDEF));
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   562
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   563
fun proofs' f = transaction0 (fn int =>
49062
7e31dfd99ce7 discontinued complicated/unreliable notion of recent proofs within context;
wenzelm
parents: 49042
diff changeset
   564
  (fn Proof (prf, x) => Proof (Proof_Node.applys (f int) prf, x)
51555
6aa64925db77 explicit Toplevel.is_skipped_proof;
wenzelm
parents: 51553
diff changeset
   565
    | skip as Skipped_Proof _ => skip
16815
13d20ed9086c added print_state_hook;
wenzelm
parents: 16729
diff changeset
   566
    | _ => raise UNDEF));
15668
53c049a365cf improved exn_message;
wenzelm
parents: 15633
diff changeset
   567
49863
b5fb6e7f8d81 more informative errors for 'proof' and 'apply' steps;
wenzelm
parents: 49062
diff changeset
   568
fun proof' f = proofs' ((Seq.single o Seq.Result) oo f);
17904
21c6894b5998 simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
wenzelm
parents: 17513
diff changeset
   569
val proofs = proofs' o K;
6689
wenzelm
parents: 6664
diff changeset
   570
val proof = proof' o K;
16815
13d20ed9086c added print_state_hook;
wenzelm
parents: 16729
diff changeset
   571
60693
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
   572
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
   573
(* skipped proofs *)
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
   574
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   575
fun actual_proof f = transaction0 (fn _ =>
21007
1bbb31aaf98d renamed enter_forward_proof to enter_proof_body;
wenzelm
parents: 20985
diff changeset
   576
  (fn Proof (prf, x) => Proof (f prf, x)
20963
a7fd8f05a2be added type global_theory -- theory or local_theory;
wenzelm
parents: 20928
diff changeset
   577
    | _ => raise UNDEF));
16815
13d20ed9086c added print_state_hook;
wenzelm
parents: 16729
diff changeset
   578
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   579
fun skip_proof f = transaction0 (fn _ =>
60693
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
   580
  (fn skip as Skipped_Proof _ => (f (); skip)
18563
wenzelm
parents: 17904
diff changeset
   581
    | _ => raise UNDEF));
wenzelm
parents: 17904
diff changeset
   582
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   583
val skip_proof_open = transaction0 (fn _ =>
60693
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
   584
  (fn Skipped_Proof (d, x) => Skipped_Proof (d + 1, x)
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
   585
    | _ => raise UNDEF));
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
   586
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   587
val skip_proof_close = transaction0 (fn _ =>
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   588
  (fn Skipped_Proof (0, (gthy, _)) => Theory gthy
60693
044f8bb3dd30 more accurate skip_proofs nesting, e.g. relevant for 'subgoal' command;
wenzelm
parents: 60403
diff changeset
   589
    | Skipped_Proof (d, x) => Skipped_Proof (d - 1, x)
33725
a8481da77270 implicit name space grouping for theory/local_theory transactions;
wenzelm
parents: 33671
diff changeset
   590
    | _ => raise UNDEF));
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   591
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   592
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   593
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   594
(** toplevel transactions **)
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   595
52527
dbac84eab3bc separate exec_id assignment for Command.print states, without affecting result of eval;
wenzelm
parents: 52499
diff changeset
   596
(* runtime position *)
27427
f6751d265cf6 added name_of;
wenzelm
parents: 26982
diff changeset
   597
52536
3a35ce87a55c tuned signature;
wenzelm
parents: 52527
diff changeset
   598
fun exec_id id (tr as Transition {pos, ...}) =
3a35ce87a55c tuned signature;
wenzelm
parents: 52527
diff changeset
   599
  position (Position.put_id (Document_ID.print id) pos) tr;
25799
7a204e0467f8 maintain thread transition properties;
wenzelm
parents: 25796
diff changeset
   600
25960
1f9956e0bd89 removed unused properties;
wenzelm
parents: 25847
diff changeset
   601
fun setmp_thread_position (Transition {pos, ...}) f x =
25819
e6feb08b7f4b replaced thread_properties by simplified version in position.ML;
wenzelm
parents: 25809
diff changeset
   602
  Position.setmp_thread_data pos f x;
25799
7a204e0467f8 maintain thread transition properties;
wenzelm
parents: 25796
diff changeset
   603
7a204e0467f8 maintain thread transition properties;
wenzelm
parents: 25796
diff changeset
   604
28095
7eaf0813bdc3 added add_hook interface for post-transition hooks;
wenzelm
parents: 27859
diff changeset
   605
(* post-transition hooks *)
7eaf0813bdc3 added add_hook interface for post-transition hooks;
wenzelm
parents: 27859
diff changeset
   606
37905
wenzelm
parents: 37865
diff changeset
   607
local
56147
9589605bcf41 prefer more robust Synchronized.var;
wenzelm
parents: 56057
diff changeset
   608
  val hooks =
9589605bcf41 prefer more robust Synchronized.var;
wenzelm
parents: 56057
diff changeset
   609
    Synchronized.var "Toplevel.hooks" ([]: (transition -> state -> state -> unit) list);
37905
wenzelm
parents: 37865
diff changeset
   610
in
28095
7eaf0813bdc3 added add_hook interface for post-transition hooks;
wenzelm
parents: 27859
diff changeset
   611
56147
9589605bcf41 prefer more robust Synchronized.var;
wenzelm
parents: 56057
diff changeset
   612
fun add_hook hook = Synchronized.change hooks (cons hook);
9589605bcf41 prefer more robust Synchronized.var;
wenzelm
parents: 56057
diff changeset
   613
fun get_hooks () = Synchronized.value hooks;
28095
7eaf0813bdc3 added add_hook interface for post-transition hooks;
wenzelm
parents: 27859
diff changeset
   614
7eaf0813bdc3 added add_hook interface for post-transition hooks;
wenzelm
parents: 27859
diff changeset
   615
end;
7eaf0813bdc3 added add_hook interface for post-transition hooks;
wenzelm
parents: 27859
diff changeset
   616
7eaf0813bdc3 added add_hook interface for post-transition hooks;
wenzelm
parents: 27859
diff changeset
   617
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   618
(* apply transitions *)
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   619
6664
f679ddd1ddd8 cleaned comments;
wenzelm
parents: 6244
diff changeset
   620
local
f679ddd1ddd8 cleaned comments;
wenzelm
parents: 6244
diff changeset
   621
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   622
fun app int (tr as Transition {markers, trans, ...}) =
67932
04352338f7f3 clarified signature;
wenzelm
parents: 67643
diff changeset
   623
  setmp_thread_position tr
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   624
    (Timing.protocol (name_of tr) (pos_of tr) (apply_trans int trans markers)
67932
04352338f7f3 clarified signature;
wenzelm
parents: 67643
diff changeset
   625
      ##> Option.map (fn UNDEF => ERROR (type_error tr) | exn => exn));
6664
f679ddd1ddd8 cleaned comments;
wenzelm
parents: 6244
diff changeset
   626
f679ddd1ddd8 cleaned comments;
wenzelm
parents: 6244
diff changeset
   627
in
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   628
26602
5534b6a6b810 made purely value-oriented, moved global state to structure Isar (cf. isar.ML);
wenzelm
parents: 26491
diff changeset
   629
fun transition int tr st =
28095
7eaf0813bdc3 added add_hook interface for post-transition hooks;
wenzelm
parents: 27859
diff changeset
   630
  let
60895
501be4aa75b4 default ML context for all command transactions, e.g. relevant for debugging and toplevel pretty-printing;
wenzelm
parents: 60693
diff changeset
   631
    val (st', opt_err) =
67642
10ff1f077119 more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents: 67641
diff changeset
   632
      Context.setmp_generic_context (try (Context.Proof o presentation_context0) st)
60895
501be4aa75b4 default ML context for all command transactions, e.g. relevant for debugging and toplevel pretty-printing;
wenzelm
parents: 60693
diff changeset
   633
        (fn () => app int tr st) ();
59055
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
   634
    val opt_err' = opt_err |> Option.map
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
   635
      (fn Runtime.EXCURSION_FAIL exn_info => exn_info
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
   636
        | exn => (Runtime.exn_context (try context_of st) exn, at_command tr));
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
   637
    val _ = get_hooks () |> List.app (fn f => (try (fn () => f tr st st') (); ()));
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
   638
  in (st', opt_err') end;
6664
f679ddd1ddd8 cleaned comments;
wenzelm
parents: 6244
diff changeset
   639
f679ddd1ddd8 cleaned comments;
wenzelm
parents: 6244
diff changeset
   640
end;
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   641
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   642
51284
59a03019f3bf fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents: 51282
diff changeset
   643
(* managed commands *)
5828
1feeadaad6a9 The Isabelle/Isar toplevel.
wenzelm
parents:
diff changeset
   644
51323
wenzelm
parents: 51322
diff changeset
   645
fun command_errors int tr st =
wenzelm
parents: 51322
diff changeset
   646
  (case transition int tr st of
59055
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
   647
    (st', NONE) => ([], SOME st')
65948
de7888573ed7 clarified build errors;
wenzelm
parents: 65059
diff changeset
   648
  | (_, SOME (exn, _)) => (Runtime.exn_messages exn, NONE));
51323
wenzelm
parents: 51322
diff changeset
   649
51284
59a03019f3bf fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents: 51282
diff changeset
   650
fun command_exception int tr st =
59a03019f3bf fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents: 51282
diff changeset
   651
  (case transition int tr st of
59055
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
   652
    (st', NONE) => st'
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
   653
  | (_, SOME (exn, info)) =>
62505
9e2a65912111 clarified modules;
wenzelm
parents: 62239
diff changeset
   654
      if Exn.is_interrupt exn then Exn.reraise exn
59055
5a7157b8e870 more informative failure of protocol commands, with exception trace;
wenzelm
parents: 59032
diff changeset
   655
      else raise Runtime.EXCURSION_FAIL (exn, info));
27576
7afff36043e6 eliminated internal command history -- superceeded by global Isar state (cf. isar.ML);
wenzelm
parents: 27564
diff changeset
   656
58848
fd0c85d7da38 eliminated odd flags and hook;
wenzelm
parents: 58798
diff changeset
   657
val command = command_exception false;
51284
59a03019f3bf fork diagnostic commands (theory loader and PIDE interaction);
wenzelm
parents: 51282
diff changeset
   658
28433
b3dab95f098f begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents: 28425
diff changeset
   659
56937
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   660
(* reset state *)
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   661
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   662
local
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   663
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   664
fun reset_state check trans st =
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   665
  if check st then NONE
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   666
  else #2 (command_errors false (trans empty) st);
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   667
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   668
in
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   669
68876
cefaac3d24ff no reset_proof for notepad: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68875
diff changeset
   670
val reset_theory = reset_state is_theory forget_proof;
56937
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   671
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   672
val reset_proof =
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   673
  reset_state is_proof
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   674
    (transaction0 (fn _ =>
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   675
      (fn Theory gthy => Skipped_Proof (0, (gthy, gthy))
56937
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   676
        | _ => raise UNDEF)));
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   677
68877
33d78e5e0a00 more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68876
diff changeset
   678
val reset_notepad =
33d78e5e0a00 more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68876
diff changeset
   679
  reset_state
33d78e5e0a00 more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68876
diff changeset
   680
    (fn st =>
33d78e5e0a00 more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68876
diff changeset
   681
      (case try proof_of st of
33d78e5e0a00 more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68876
diff changeset
   682
        SOME state => not (Proof.is_notepad state) orelse can Proof.end_notepad state
33d78e5e0a00 more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68876
diff changeset
   683
      | NONE => true))
68878
9203eb13bef7 clarified reset_notepad;
wenzelm
parents: 68877
diff changeset
   684
    (proof Proof.reset_notepad);
68877
33d78e5e0a00 more robust reset_state: begin/end structure takes precedence over goal/proof structure;
wenzelm
parents: 68876
diff changeset
   685
56937
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   686
end;
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   687
d11d11da0d90 smarter recovery from toplevel type error;
wenzelm
parents: 56897
diff changeset
   688
46959
cdc791910460 defer actual parsing of command spans and thus allow new commands to be used in the same theory where defined;
wenzelm
parents: 45666
diff changeset
   689
(* scheduled proof result *)
28433
b3dab95f098f begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents: 28425
diff changeset
   690
51332
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   691
datatype result =
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   692
  Result of transition * state |
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   693
  Result_List of result list |
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   694
  Result_Future of result future;
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   695
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   696
fun join_results (Result x) = [x]
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   697
  | join_results (Result_List xs) = maps join_results xs
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   698
  | join_results (Result_Future x) = join_results (Future.join x);
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   699
51323
wenzelm
parents: 51322
diff changeset
   700
local
wenzelm
parents: 51322
diff changeset
   701
47417
9679bab23f93 misc tuning and simplification;
wenzelm
parents: 47416
diff changeset
   702
structure Result = Proof_Data
28974
d6b190efa01a excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents: 28645
diff changeset
   703
(
51332
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   704
  type T = result;
59150
wenzelm
parents: 59055
diff changeset
   705
  fun init _ = Result_List [];
28974
d6b190efa01a excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents: 28645
diff changeset
   706
);
d6b190efa01a excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents: 28645
diff changeset
   707
51332
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   708
val get_result = Result.get o Proof.context_of;
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   709
val put_result = Proof.map_context o Result.put;
51324
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   710
66169
8cfa8c7ee1f6 keep original bottom-up order of proof forks, which potentially reduces thread congestion due to Proofterm.consolidate;
wenzelm
parents: 65948
diff changeset
   711
fun timing_estimate elem =
68839
d8251a61cce8 clarified modules;
wenzelm
parents: 68772
diff changeset
   712
  let val trs = tl (Thy_Element.flat_element elem)
69877
45b2e784350a tuned signature;
wenzelm
parents: 69719
diff changeset
   713
  in fold (fn tr => fn t => timing_of tr + t) trs Time.zeroTime end;
51423
e5f9a6d9ca82 clarified parallel_subproofs_saturation (blind guess) vs. parallel_subproofs_threshold (precient timing estimate);
wenzelm
parents: 51332
diff changeset
   714
68130
6fb85346cb79 clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents: 67932
diff changeset
   715
fun future_proofs_enabled estimate st =
51324
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   716
  (case try proof_of st of
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   717
    NONE => false
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   718
  | SOME state =>
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   719
      not (Proof.is_relevant state) andalso
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   720
       (if can (Proof.assert_bottom true) state
68130
6fb85346cb79 clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents: 67932
diff changeset
   721
        then Future.proofs_enabled 1
6fb85346cb79 clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents: 67932
diff changeset
   722
        else Future.proofs_enabled 2 orelse Future.proofs_enabled_timing estimate));
51278
wenzelm
parents: 51274
diff changeset
   723
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   724
val empty_markers = markers [];
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   725
val empty_trans = reset_trans #> keep (K ());
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   726
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   727
in
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   728
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   729
fun fork_presentation tr = (tr |> empty_markers, tr |> empty_trans);
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   730
58923
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58893
diff changeset
   731
fun atom_result keywords tr st =
51323
wenzelm
parents: 51322
diff changeset
   732
  let
wenzelm
parents: 51322
diff changeset
   733
    val st' =
68130
6fb85346cb79 clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents: 67932
diff changeset
   734
      if Future.proofs_enabled 1 andalso Keyword.is_diag keywords (name_of tr) then
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   735
        let
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   736
          val (tr1, tr2) = fork_presentation tr;
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   737
          val _ =
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   738
            Execution.fork {name = "Toplevel.diag", pos = pos_of tr, pri = ~1}
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   739
              (fn () => command tr1 st);
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   740
        in command tr2 st end
51323
wenzelm
parents: 51322
diff changeset
   741
      else command tr st;
51332
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   742
  in (Result (tr, st'), st') end;
51323
wenzelm
parents: 51322
diff changeset
   743
68839
d8251a61cce8 clarified modules;
wenzelm
parents: 68772
diff changeset
   744
fun element_result keywords (Thy_Element.Element (tr, NONE)) st = atom_result keywords tr st
d8251a61cce8 clarified modules;
wenzelm
parents: 68772
diff changeset
   745
  | element_result keywords (elem as Thy_Element.Element (head_tr, SOME element_rest)) st =
48633
7cd32f9d4293 recovered comination of Toplevel.skip_proofs and Goal.parallel_proofs from 9679bab23f93 (NB: skip_proofs leads to failure of Toplevel.proof_of);
wenzelm
parents: 47881
diff changeset
   746
      let
58923
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58893
diff changeset
   747
        val (head_result, st') = atom_result keywords head_tr st;
51332
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   748
        val (body_elems, end_tr) = element_rest;
66169
8cfa8c7ee1f6 keep original bottom-up order of proof forks, which potentially reduces thread congestion due to Proofterm.consolidate;
wenzelm
parents: 65948
diff changeset
   749
        val estimate = timing_estimate elem;
51324
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   750
      in
68130
6fb85346cb79 clarified future scheduling parameters, with support for parallel_limit;
wenzelm
parents: 67932
diff changeset
   751
        if not (future_proofs_enabled estimate st')
51324
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   752
        then
51332
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   753
          let
68839
d8251a61cce8 clarified modules;
wenzelm
parents: 68772
diff changeset
   754
            val proof_trs = maps Thy_Element.flat_element body_elems @ [end_tr];
58923
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58893
diff changeset
   755
            val (proof_results, st'') = fold_map (atom_result keywords) proof_trs st';
51332
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   756
          in (Result_List (head_result :: proof_results), st'') end
51324
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   757
        else
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   758
          let
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   759
            val (end_tr1, end_tr2) = fork_presentation end_tr;
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   760
51324
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   761
            val finish = Context.Theory o Proof_Context.theory_of;
28974
d6b190efa01a excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents: 28645
diff changeset
   762
51605
eca8acb42e4a more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents: 51595
diff changeset
   763
            val future_proof =
eca8acb42e4a more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents: 51595
diff changeset
   764
              Proof.future_proof (fn state =>
53192
04df1d236e1c maintain goal forks as part of global execution;
wenzelm
parents: 53189
diff changeset
   765
                Execution.fork
66169
8cfa8c7ee1f6 keep original bottom-up order of proof forks, which potentially reduces thread congestion due to Proofterm.consolidate;
wenzelm
parents: 65948
diff changeset
   766
                  {name = "Toplevel.future_proof", pos = pos_of head_tr, pri = ~1}
51605
eca8acb42e4a more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents: 51595
diff changeset
   767
                  (fn () =>
eca8acb42e4a more explicit Goal.fork_params -- avoid implicit arguments via thread data;
wenzelm
parents: 51595
diff changeset
   768
                    let
69883
f8293bf510a0 clarified Toplevel.state: more explicit types;
wenzelm
parents: 69882
diff changeset
   769
                      val State ((Proof (prf, (_, orig_gthy)), _), prev_thy) = st';
69878
ccc8e4c99520 tuned -- more explicit type node_presentation;
wenzelm
parents: 69877
diff changeset
   770
                      val node' = Proof (Proof_Node.apply (K state) prf, (finish, orig_gthy));
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   771
                      val (results, result_state) =
69886
0cb8753bdb50 clarified signature;
wenzelm
parents: 69883
diff changeset
   772
                        State (node_presentation node', prev_thy)
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   773
                        |> fold_map (element_result keywords) body_elems ||> command end_tr1;
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   774
                    in (Result_List results, presentation_context0 result_state) end))
51332
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   775
              #> (fn (res, state') => state' |> put_result (Result_Future res));
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   776
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   777
            val forked_proof =
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   778
              proof (future_proof #>
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   779
                (fn state => state |> Proof.local_done_proof |> put_result (get_result state))) o
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   780
              end_proof (fn _ => future_proof #>
8707df0b0255 refined parallel_proofs = 2: fork whole Isar sub-proofs, not just terminal ones;
wenzelm
parents: 51324
diff changeset
   781
                (fn state => state |> Proof.global_done_proof |> Result.put (get_result state)));
28974
d6b190efa01a excursion: pass explicit proof states as result of future proof, replaced low-level Thm.join_futures by PureThy.force_proofs;
wenzelm
parents: 28645
diff changeset
   782
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   783
            val st'' = st' |> command (head_tr |> reset_trans |> forked_proof);
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   784
            val end_st = st'' |> command end_tr2;
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   785
            val end_result = Result (end_tr, end_st);
51324
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   786
            val result =
67642
10ff1f077119 more tight presentation context: avoid storing full Toplevel.state;
wenzelm
parents: 67641
diff changeset
   787
              Result_List [head_result, Result.get (presentation_context0 st''), end_result];
69887
b9985133805d added semantic document markers;
wenzelm
parents: 69886
diff changeset
   788
          in (result, end_st) end
51324
c17f5de0a915 uniform treatment of global/local proofs;
wenzelm
parents: 51323
diff changeset
   789
      end;
28433
b3dab95f098f begin_proof: avoid race condition wrt. skip_proofs flag;
wenzelm
parents: 28425
diff changeset
   790
6664
f679ddd1ddd8 cleaned comments;
wenzelm
parents: 6244
diff changeset
   791
end;
51323
wenzelm
parents: 51322
diff changeset
   792
wenzelm
parents: 51322
diff changeset
   793
end;
69719
331ef175a112 Backed out changeset 1bc422c08209 -- obsolete in AFP/5d11846ac6ab;
wenzelm
parents: 69715
diff changeset
   794
69708
1c201e4792cb Local_Theory.reset only required for toplevel interaction, attempt to withhold it from user space
haftmann
parents: 68878
diff changeset
   795
structure Local_Theory : LOCAL_THEORY = struct open Local_Theory; end;