src/Pure/Isar/toplevel.ML
author wenzelm
Mon, 01 Oct 2007 15:14:57 +0200
changeset 24795 6f5cb7885fd7
parent 24780 47bb1e380d83
child 24920 2a45e400fdad
permissions -rw-r--r--
print_state_context: local theory context, not proof context;
context_position: cover Theory case as well (requires additional checkpoint);
wenzelm@5828
     1
(*  Title:      Pure/Isar/toplevel.ML
wenzelm@5828
     2
    ID:         $Id$
wenzelm@5828
     3
    Author:     Markus Wenzel, TU Muenchen
wenzelm@5828
     4
wenzelm@5828
     5
The Isabelle/Isar toplevel.
wenzelm@5828
     6
*)
wenzelm@5828
     7
wenzelm@5828
     8
signature TOPLEVEL =
wenzelm@5828
     9
sig
wenzelm@19063
    10
  exception UNDEF
wenzelm@20963
    11
  type generic_theory
wenzelm@18589
    12
  type node
wenzelm@20963
    13
  val theory_node: node -> generic_theory option
wenzelm@18589
    14
  val proof_node: node -> ProofHistory.T option
wenzelm@20963
    15
  val cases_node: (generic_theory -> 'a) -> (Proof.state -> 'a) -> node -> 'a
wenzelm@20963
    16
  val presentation_context: node option -> xstring option -> Proof.context
wenzelm@5828
    17
  type state
wenzelm@7732
    18
  val is_toplevel: state -> bool
wenzelm@18589
    19
  val is_theory: state -> bool
wenzelm@18589
    20
  val is_proof: state -> bool
wenzelm@17076
    21
  val level: state -> int
wenzelm@6689
    22
  val node_history_of: state -> node History.T
wenzelm@5828
    23
  val node_of: state -> node
wenzelm@20963
    24
  val node_case: (generic_theory -> 'a) -> (Proof.state -> 'a) -> state -> 'a
wenzelm@21506
    25
  val context_of: state -> Proof.context
wenzelm@22089
    26
  val generic_theory_of: state -> generic_theory
wenzelm@5828
    27
  val theory_of: state -> theory
wenzelm@5828
    28
  val proof_of: state -> Proof.state
wenzelm@18589
    29
  val proof_position_of: state -> int
wenzelm@21007
    30
  val enter_proof_body: state -> Proof.state
wenzelm@16815
    31
  val print_state_context: state -> unit
wenzelm@16815
    32
  val print_state: bool -> state -> unit
wenzelm@16682
    33
  val quiet: bool ref
wenzelm@16682
    34
  val debug: bool ref
wenzelm@17321
    35
  val interact: bool ref
wenzelm@16682
    36
  val timing: bool ref
wenzelm@16682
    37
  val profiling: int ref
wenzelm@16815
    38
  val skip_proofs: bool ref
wenzelm@5828
    39
  exception TERMINATE
wenzelm@5990
    40
  exception RESTART
wenzelm@24055
    41
  exception TOPLEVEL_ERROR
wenzelm@20128
    42
  val exn_message: exn -> string
wenzelm@20128
    43
  val program: (unit -> 'a) -> 'a
wenzelm@16682
    44
  type transition
wenzelm@6689
    45
  val undo_limit: bool -> int option
wenzelm@5828
    46
  val empty: transition
wenzelm@14923
    47
  val name_of: transition -> string
wenzelm@14923
    48
  val source_of: transition -> OuterLex.token list option
wenzelm@5828
    49
  val name: string -> transition -> transition
wenzelm@5828
    50
  val position: Position.T -> transition -> transition
wenzelm@14923
    51
  val source: OuterLex.token list -> transition -> transition
wenzelm@5828
    52
  val interactive: bool -> transition -> transition
wenzelm@5828
    53
  val print: transition -> transition
wenzelm@16607
    54
  val print': string -> transition -> transition
wenzelm@17363
    55
  val three_buffersN: string
wenzelm@17363
    56
  val print3: transition -> transition
wenzelm@9010
    57
  val no_timing: transition -> transition
wenzelm@22056
    58
  val init_theory: (bool -> theory) -> (theory -> unit) -> (theory -> unit) ->
wenzelm@22056
    59
    transition -> transition
wenzelm@23911
    60
  val init_empty: (unit -> unit) -> transition -> transition
wenzelm@6689
    61
  val exit: transition -> transition
wenzelm@21958
    62
  val undo_exit: transition -> transition
wenzelm@6689
    63
  val kill: transition -> transition
wenzelm@20128
    64
  val history: (node History.T -> node History.T) -> transition -> transition
wenzelm@5828
    65
  val keep: (state -> unit) -> transition -> transition
wenzelm@7612
    66
  val keep': (bool -> state -> unit) -> transition -> transition
wenzelm@5828
    67
  val imperative: (unit -> unit) -> transition -> transition
wenzelm@5828
    68
  val theory: (theory -> theory) -> transition -> transition
wenzelm@7612
    69
  val theory': (bool -> theory -> theory) -> transition -> transition
wenzelm@20985
    70
  val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition
wenzelm@21007
    71
  val end_local_theory: transition -> transition
wenzelm@20963
    72
  val local_theory: xstring option -> (local_theory -> local_theory) -> transition -> transition
wenzelm@20963
    73
  val present_local_theory: xstring option -> (bool -> node -> unit) -> transition -> transition
berghofe@24453
    74
  val local_theory_to_proof': xstring option -> (bool -> local_theory -> Proof.state) ->
berghofe@24453
    75
    transition -> transition
wenzelm@21007
    76
  val local_theory_to_proof: xstring option -> (local_theory -> Proof.state) ->
wenzelm@21007
    77
    transition -> transition
wenzelm@17363
    78
  val theory_to_proof: (theory -> Proof.state) -> transition -> transition
wenzelm@21007
    79
  val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition
wenzelm@21007
    80
  val forget_proof: transition -> transition
wenzelm@21177
    81
  val present_proof: (bool -> node -> unit) -> transition -> transition
wenzelm@21177
    82
  val proofs': (bool -> Proof.state -> Proof.state Seq.seq) -> transition -> transition
wenzelm@21177
    83
  val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition
wenzelm@21177
    84
  val proofs: (Proof.state -> Proof.state Seq.seq) -> transition -> transition
wenzelm@17904
    85
  val proof: (Proof.state -> Proof.state) -> transition -> transition
wenzelm@16815
    86
  val actual_proof: (ProofHistory.T -> ProofHistory.T) -> transition -> transition
wenzelm@16815
    87
  val skip_proof: (int History.T -> int History.T) -> transition -> transition
wenzelm@17904
    88
  val skip_proof_to_theory: (int -> bool) -> transition -> transition
wenzelm@9512
    89
  val unknown_theory: transition -> transition
wenzelm@9512
    90
  val unknown_proof: transition -> transition
wenzelm@9512
    91
  val unknown_context: transition -> transition
wenzelm@17076
    92
  val present_excursion: (transition * (state -> state -> 'a -> 'a)) list -> 'a -> 'a
wenzelm@5828
    93
  val excursion: transition list -> unit
wenzelm@5828
    94
  val set_state: state -> unit
wenzelm@5828
    95
  val get_state: unit -> state
wenzelm@5828
    96
  val exn: unit -> (exn * string) option
wenzelm@5828
    97
  val >> : transition -> bool
wenzelm@14985
    98
  val >>> : transition list -> unit
wenzelm@21958
    99
  val init_state: unit -> unit
nipkow@14091
   100
  type 'a isar
nipkow@14091
   101
  val loop: 'a isar -> unit
wenzelm@5828
   102
end;
wenzelm@5828
   103
wenzelm@6965
   104
structure Toplevel: TOPLEVEL =
wenzelm@5828
   105
struct
wenzelm@5828
   106
wenzelm@21958
   107
wenzelm@5828
   108
(** toplevel state **)
wenzelm@5828
   109
wenzelm@19063
   110
exception UNDEF;
wenzelm@19063
   111
wenzelm@19063
   112
wenzelm@21294
   113
(* local theory wrappers *)
wenzelm@5828
   114
wenzelm@20963
   115
type generic_theory = Context.generic;    (*theory or local_theory*)
wenzelm@20963
   116
wenzelm@21294
   117
val loc_init = TheoryTarget.init;
wenzelm@21294
   118
wenzelm@21294
   119
val loc_exit = ProofContext.theory_of o LocalTheory.exit;
wenzelm@21294
   120
wenzelm@21294
   121
fun loc_begin loc (Context.Theory thy) = loc_init loc thy
wenzelm@21294
   122
  | loc_begin NONE (Context.Proof lthy) = lthy
wenzelm@21294
   123
  | loc_begin loc (Context.Proof lthy) = loc_init loc (loc_exit lthy);
wenzelm@21294
   124
wenzelm@21294
   125
fun loc_finish _ (Context.Theory _) = Context.Theory o loc_exit
wenzelm@21294
   126
  | loc_finish NONE (Context.Proof _) = Context.Proof o LocalTheory.restore
wenzelm@21294
   127
  | loc_finish (SOME _) (Context.Proof lthy) = Context.Proof o LocalTheory.reinit lthy o loc_exit;
wenzelm@21294
   128
wenzelm@21294
   129
wenzelm@21958
   130
(* datatype node *)
wenzelm@21294
   131
wenzelm@5828
   132
datatype node =
wenzelm@20963
   133
  Theory of generic_theory * Proof.context option | (*theory with presentation context*)
wenzelm@21007
   134
  Proof of ProofHistory.T * ((Proof.context -> generic_theory) * generic_theory) |
wenzelm@21007
   135
    (*history of proof states, finish, original theory*)
wenzelm@21007
   136
  SkipProof of int History.T * (generic_theory * generic_theory);
wenzelm@18563
   137
    (*history of proof depths, resulting theory, original theory*)
wenzelm@5828
   138
wenzelm@22056
   139
val the_global_theory = fn Theory (Context.Theory thy, _) => thy | _ => raise UNDEF;
wenzelm@20963
   140
val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE;
wenzelm@18589
   141
val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE;
wenzelm@18589
   142
wenzelm@20963
   143
fun cases_node f _ (Theory (gthy, _)) = f gthy
wenzelm@19063
   144
  | cases_node _ g (Proof (prf, _)) = g (ProofHistory.current prf)
wenzelm@21007
   145
  | cases_node f _ (SkipProof (_, (gthy, _))) = f gthy;
wenzelm@19063
   146
wenzelm@20963
   147
fun presentation_context (SOME (Theory (_, SOME ctxt))) NONE = ctxt
wenzelm@20963
   148
  | presentation_context (SOME node) NONE = cases_node Context.proof_of Proof.context_of node
wenzelm@20963
   149
  | presentation_context (SOME node) (SOME loc) =
wenzelm@21294
   150
      loc_init (SOME loc) (cases_node Context.theory_of Proof.theory_of node)
wenzelm@20963
   151
  | presentation_context NONE _ = raise UNDEF;
wenzelm@19063
   152
wenzelm@5828
   153
wenzelm@21958
   154
(* datatype state *)
wenzelm@5828
   155
wenzelm@22056
   156
type state_info = node History.T * ((theory -> unit) * (theory -> unit));
wenzelm@21958
   157
wenzelm@21958
   158
datatype state =
wenzelm@21958
   159
  Toplevel of state_info option |  (*outer toplevel, leftover end state*)
wenzelm@21958
   160
  State of state_info;
wenzelm@21958
   161
wenzelm@21958
   162
val toplevel = Toplevel NONE;
wenzelm@21958
   163
wenzelm@21958
   164
fun is_toplevel (Toplevel _) = true
wenzelm@7732
   165
  | is_toplevel _ = false;
wenzelm@7732
   166
wenzelm@21958
   167
fun level (Toplevel _) = 0
wenzelm@21958
   168
  | level (State (node, _)) =
wenzelm@17076
   169
      (case History.current node of
wenzelm@21310
   170
        Theory _ => 0
wenzelm@21310
   171
      | Proof (prf, _) => Proof.level (ProofHistory.current prf)
wenzelm@21310
   172
      | SkipProof (h, _) => History.current h + 1);   (*different notion of proof depth!*)
wenzelm@17076
   173
wenzelm@21958
   174
fun str_of_state (Toplevel _) = "at top level"
wenzelm@21958
   175
  | str_of_state (State (node, _)) =
wenzelm@16815
   176
      (case History.current node of
wenzelm@20963
   177
        Theory (Context.Theory _, _) => "in theory mode"
wenzelm@20963
   178
      | Theory (Context.Proof _, _) => "in local theory mode"
wenzelm@16815
   179
      | Proof _ => "in proof mode"
wenzelm@16815
   180
      | SkipProof _ => "in skipped proof mode");
wenzelm@5946
   181
wenzelm@5946
   182
wenzelm@5828
   183
(* top node *)
wenzelm@5828
   184
wenzelm@21958
   185
fun node_history_of (Toplevel _) = raise UNDEF
wenzelm@21958
   186
  | node_history_of (State (node, _)) = node;
wenzelm@6689
   187
wenzelm@6689
   188
val node_of = History.current o node_history_of;
wenzelm@5828
   189
wenzelm@18589
   190
fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state));
wenzelm@18589
   191
fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state));
wenzelm@18589
   192
wenzelm@19063
   193
fun node_case f g state = cases_node f g (node_of state);
wenzelm@5828
   194
wenzelm@21506
   195
val context_of = node_case Context.proof_of Proof.context_of;
wenzelm@22089
   196
val generic_theory_of = node_case I (Context.Proof o Proof.context_of);
wenzelm@20963
   197
val theory_of = node_case Context.theory_of Proof.theory_of;
wenzelm@18589
   198
val proof_of = node_case (fn _ => raise UNDEF) I;
wenzelm@17208
   199
wenzelm@18589
   200
fun proof_position_of state =
wenzelm@18589
   201
  (case node_of state of
wenzelm@18589
   202
    Proof (prf, _) => ProofHistory.position prf
wenzelm@18589
   203
  | _ => raise UNDEF);
wenzelm@6664
   204
wenzelm@21007
   205
val enter_proof_body = node_case (Proof.init o Context.proof_of) Proof.enter_forward;
wenzelm@5828
   206
wenzelm@5828
   207
wenzelm@16815
   208
(* print state *)
wenzelm@16815
   209
wenzelm@21294
   210
val pretty_context = LocalTheory.pretty o Context.cases (loc_init NONE) I;
wenzelm@16815
   211
wenzelm@23640
   212
fun print_state_context state =
wenzelm@24795
   213
  (case try node_of state of
wenzelm@21506
   214
    NONE => []
wenzelm@24795
   215
  | SOME (Theory (gthy, _)) => pretty_context gthy
wenzelm@24795
   216
  | SOME (Proof (_, (_, gthy))) => pretty_context gthy
wenzelm@24795
   217
  | SOME (SkipProof (_, (gthy, _))) => pretty_context gthy)
wenzelm@23640
   218
  |> Pretty.chunks |> Pretty.writeln;
wenzelm@16815
   219
wenzelm@23640
   220
fun print_state prf_only state =
wenzelm@23701
   221
  (case try node_of state of
wenzelm@23701
   222
    NONE => []
wenzelm@23701
   223
  | SOME (Theory (gthy, _)) => if prf_only then [] else pretty_context gthy
wenzelm@23701
   224
  | SOME (Proof (prf, _)) =>
wenzelm@23701
   225
      Proof.pretty_state (ProofHistory.position prf) (ProofHistory.current prf)
wenzelm@23701
   226
  | SOME (SkipProof (h, _)) =>
wenzelm@23701
   227
      [Pretty.str ("skipped proof: depth " ^ string_of_int (History.current h))])
wenzelm@23701
   228
  |> Pretty.markup_chunks Markup.state |> Pretty.writeln;
wenzelm@16815
   229
wenzelm@16815
   230
wenzelm@15668
   231
wenzelm@5828
   232
(** toplevel transitions **)
wenzelm@5828
   233
wenzelm@16682
   234
val quiet = ref false;
wenzelm@22135
   235
val debug = Output.debugging;
wenzelm@17321
   236
val interact = ref false;
wenzelm@16682
   237
val timing = Output.timing;
wenzelm@16682
   238
val profiling = ref 0;
wenzelm@16815
   239
val skip_proofs = ref false;
wenzelm@16682
   240
wenzelm@5828
   241
exception TERMINATE;
wenzelm@5990
   242
exception RESTART;
wenzelm@7022
   243
exception EXCURSION_FAIL of exn * string;
wenzelm@6689
   244
exception FAILURE of state * exn;
wenzelm@24055
   245
exception TOPLEVEL_ERROR;
wenzelm@5828
   246
wenzelm@20128
   247
wenzelm@20128
   248
(* print exceptions *)
wenzelm@20128
   249
wenzelm@20128
   250
local
wenzelm@20128
   251
wenzelm@20128
   252
fun with_context f xs =
wenzelm@22095
   253
  (case ML_Context.get_context () of NONE => []
wenzelm@22089
   254
  | SOME context => map (f (Context.proof_of context)) xs);
wenzelm@20128
   255
wenzelm@20128
   256
fun raised name [] = "exception " ^ name ^ " raised"
wenzelm@20128
   257
  | raised name [msg] = "exception " ^ name ^ " raised: " ^ msg
wenzelm@20128
   258
  | raised name msgs = cat_lines (("exception " ^ name ^ " raised:") :: msgs);
wenzelm@20128
   259
wenzelm@20128
   260
fun exn_msg _ TERMINATE = "Exit."
wenzelm@20128
   261
  | exn_msg _ RESTART = "Restart."
wenzelm@20128
   262
  | exn_msg _ Interrupt = "Interrupt."
wenzelm@24055
   263
  | exn_msg _ TOPLEVEL_ERROR = "Error."
wenzelm@20128
   264
  | exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg
wenzelm@20128
   265
  | exn_msg _ (ERROR msg) = msg
wenzelm@23975
   266
  | exn_msg detailed (Exn.EXCEPTIONS (exns, "")) = cat_lines (map (exn_msg detailed) exns)
wenzelm@23963
   267
  | exn_msg detailed (Exn.EXCEPTIONS (exns, msg)) = cat_lines (map (exn_msg detailed) exns @ [msg])
wenzelm@20128
   268
  | exn_msg detailed (EXCURSION_FAIL (exn, msg)) = cat_lines [exn_msg detailed exn, msg]
wenzelm@20128
   269
  | exn_msg false (THEORY (msg, _)) = msg
wenzelm@20128
   270
  | exn_msg true (THEORY (msg, thys)) = raised "THEORY" (msg :: map Context.str_of_thy thys)
wenzelm@20128
   271
  | exn_msg false (Syntax.AST (msg, _)) = raised "AST" [msg]
wenzelm@20128
   272
  | exn_msg true (Syntax.AST (msg, asts)) =
wenzelm@20128
   273
      raised "AST" (msg :: map (Pretty.string_of o Syntax.pretty_ast) asts)
wenzelm@20128
   274
  | exn_msg false (TYPE (msg, _, _)) = raised "TYPE" [msg]
wenzelm@20128
   275
  | exn_msg true (TYPE (msg, Ts, ts)) = raised "TYPE" (msg ::
wenzelm@22089
   276
        with_context ProofContext.string_of_typ Ts @ with_context ProofContext.string_of_term ts)
wenzelm@20128
   277
  | exn_msg false (TERM (msg, _)) = raised "TERM" [msg]
wenzelm@22089
   278
  | exn_msg true (TERM (msg, ts)) =
wenzelm@22089
   279
      raised "TERM" (msg :: with_context ProofContext.string_of_term ts)
wenzelm@20128
   280
  | exn_msg false (THM (msg, _, _)) = raised "THM" [msg]
wenzelm@20128
   281
  | exn_msg true (THM (msg, i, thms)) =
wenzelm@22089
   282
      raised ("THM " ^ string_of_int i) (msg :: with_context ProofContext.string_of_thm thms)
wenzelm@20128
   283
  | exn_msg _ Option.Option = raised "Option" []
wenzelm@20128
   284
  | exn_msg _ Library.UnequalLengths = raised "UnequalLengths" []
wenzelm@20128
   285
  | exn_msg _ Empty = raised "Empty" []
wenzelm@20128
   286
  | exn_msg _ Subscript = raised "Subscript" []
wenzelm@20128
   287
  | exn_msg _ (Fail msg) = raised "Fail" [msg]
wenzelm@23940
   288
  | exn_msg _ exn = General.exnMessage exn;
wenzelm@20128
   289
wenzelm@20128
   290
in
wenzelm@20128
   291
wenzelm@20128
   292
fun exn_message exn = exn_msg (! debug) exn;
wenzelm@20128
   293
wenzelm@22014
   294
fun print_exn NONE = ()
wenzelm@22014
   295
  | print_exn (SOME (exn, s)) = Output.error_msg (cat_lines [exn_message exn, s]);
wenzelm@20128
   296
wenzelm@20128
   297
end;
wenzelm@20128
   298
wenzelm@20128
   299
wenzelm@20128
   300
(* controlled execution *)
wenzelm@20128
   301
wenzelm@20128
   302
local
wenzelm@20128
   303
wenzelm@18685
   304
fun debugging f x =
wenzelm@23940
   305
  if ! debug then exception_trace (fn () => f x)
wenzelm@18685
   306
  else f x;
wenzelm@18685
   307
wenzelm@20128
   308
fun interruptible f x =
wenzelm@20128
   309
  let val y = ref x
wenzelm@20128
   310
  in raise_interrupt (fn () => y := f x) (); ! y end;
wenzelm@20128
   311
wenzelm@24055
   312
fun toplevel_error f x = f x
wenzelm@24055
   313
  handle exn => (Output.error_msg (exn_message exn); raise TOPLEVEL_ERROR);
wenzelm@24055
   314
wenzelm@20128
   315
in
wenzelm@20128
   316
wenzelm@20128
   317
fun controlled_execution f =
wenzelm@20128
   318
  f
wenzelm@20128
   319
  |> debugging
wenzelm@24055
   320
  |> interruptible;
wenzelm@20128
   321
wenzelm@20128
   322
fun program f =
wenzelm@24055
   323
 (f
wenzelm@24055
   324
  |> debugging
wenzelm@24055
   325
  |> toplevel_error) ();
wenzelm@20128
   326
wenzelm@20128
   327
end;
wenzelm@20128
   328
wenzelm@5828
   329
wenzelm@16815
   330
(* node transactions and recovery from stale theories *)
wenzelm@6689
   331
wenzelm@16815
   332
(*NB: proof commands should be non-destructive!*)
wenzelm@7022
   333
wenzelm@6689
   334
local
wenzelm@6689
   335
wenzelm@16452
   336
fun is_stale state = Context.is_stale (theory_of state) handle UNDEF => false;
wenzelm@6689
   337
wenzelm@18685
   338
val stale_theory = ERROR "Stale theory encountered after succesful execution!";
wenzelm@16815
   339
wenzelm@23363
   340
fun map_theory f = History.map_current
wenzelm@21177
   341
  (fn Theory (gthy, _) => Theory (Context.mapping f (LocalTheory.raw_theory f) gthy, NONE)
wenzelm@21177
   342
    | node => node);
wenzelm@6689
   343
wenzelm@23363
   344
fun context_position pos = History.map_current
wenzelm@24795
   345
  (fn Theory (gthy, ctxt) => Theory (ContextPosition.put pos gthy, ctxt)
wenzelm@23363
   346
    | Proof (prf, x) =>
wenzelm@24795
   347
        Proof (ProofHistory.map_current (Proof.map_context (ContextPosition.put_ctxt pos)) prf, x)
wenzelm@23363
   348
    | node => node);
wenzelm@23363
   349
skalberg@15531
   350
fun return (result, NONE) = result
skalberg@15531
   351
  | return (result, SOME exn) = raise FAILURE (result, exn);
wenzelm@7022
   352
wenzelm@6689
   353
in
wenzelm@6689
   354
wenzelm@23363
   355
fun transaction hist pos f (node, term) =
wenzelm@20128
   356
  let
wenzelm@21177
   357
    val cont_node = map_theory Theory.checkpoint node;
wenzelm@21177
   358
    val back_node = map_theory Theory.copy cont_node;
wenzelm@21958
   359
    fun state nd = State (nd, term);
wenzelm@20128
   360
    fun normal_state nd = (state nd, NONE);
wenzelm@20128
   361
    fun error_state nd exn = (state nd, SOME exn);
wenzelm@6689
   362
wenzelm@20128
   363
    val (result, err) =
wenzelm@20128
   364
      cont_node
wenzelm@23363
   365
      |> context_position pos
wenzelm@24795
   366
      |> map_theory Theory.checkpoint
wenzelm@20128
   367
      |> (f
wenzelm@23363
   368
          |> (if hist then History.apply' (History.current back_node) else History.map_current)
wenzelm@20128
   369
          |> controlled_execution)
wenzelm@23363
   370
      |> context_position Position.none
wenzelm@20128
   371
      |> normal_state
wenzelm@20128
   372
      handle exn => error_state cont_node exn;
wenzelm@20128
   373
  in
wenzelm@20128
   374
    if is_stale result
wenzelm@20128
   375
    then return (error_state back_node (the_default stale_theory err))
wenzelm@20128
   376
    else return (result, err)
wenzelm@20128
   377
  end;
wenzelm@6689
   378
wenzelm@6689
   379
end;
wenzelm@6689
   380
wenzelm@6689
   381
wenzelm@6689
   382
(* primitive transitions *)
wenzelm@6689
   383
wenzelm@18563
   384
(*Note: Recovery from stale theories is provided only for theory-level
wenzelm@18589
   385
  operations via Transaction.  Other node or state operations should
wenzelm@18589
   386
  not touch theories at all.  Interrupts are enabled only for Keep and
wenzelm@18589
   387
  Transaction.*)
wenzelm@5828
   388
wenzelm@5828
   389
datatype trans =
wenzelm@22056
   390
  Init of (bool -> theory) * ((theory -> unit) * (theory -> unit)) |
wenzelm@21958
   391
                                                    (*init node; with exit/kill operation*)
wenzelm@23911
   392
  InitEmpty of unit -> unit |                       (*init empty toplevel*)
wenzelm@21958
   393
  Exit |                                            (*conclude node -- deferred until init*)
wenzelm@21958
   394
  UndoExit |                                        (*continue after conclusion*)
wenzelm@21958
   395
  Kill |                                            (*abort node*)
wenzelm@21958
   396
  History of node History.T -> node History.T |     (*history operation (undo etc.)*)
wenzelm@21958
   397
  Keep of bool -> state -> unit |                   (*peek at state*)
wenzelm@21958
   398
  Transaction of bool * (bool -> node -> node);     (*node transaction*)
wenzelm@5828
   399
skalberg@15531
   400
fun undo_limit int = if int then NONE else SOME 0;
wenzelm@5828
   401
wenzelm@22056
   402
fun safe_exit (Toplevel (SOME (node, (exit, _)))) =
wenzelm@22056
   403
    (case try the_global_theory (History.current node) of
wenzelm@22056
   404
      SOME thy => exit thy
wenzelm@22056
   405
    | NONE => ())
wenzelm@22056
   406
  | safe_exit _ = ();
wenzelm@21958
   407
wenzelm@6689
   408
local
wenzelm@6689
   409
wenzelm@21958
   410
fun keep_state int f = controlled_execution (fn x => tap (f int) x);
wenzelm@21958
   411
wenzelm@23363
   412
fun apply_tr int _ (Init (f, term)) (state as Toplevel _) =
wenzelm@22056
   413
      let val node = Theory (Context.Theory (f int), NONE)
wenzelm@22056
   414
      in safe_exit state; State (History.init (undo_limit int) node, term) end
wenzelm@23911
   415
  | apply_tr int _ (InitEmpty f) (state as Toplevel _) =
wenzelm@23911
   416
      (safe_exit state; keep_state int (fn _ => fn _ => f ()) toplevel)
wenzelm@23363
   417
  | apply_tr _ _ Exit (State (node, term)) =
wenzelm@22056
   418
      (the_global_theory (History.current node); Toplevel (SOME (node, term)))
wenzelm@23363
   419
  | apply_tr _ _ UndoExit (Toplevel (SOME state_info)) = State state_info
wenzelm@23363
   420
  | apply_tr _ _ Kill (State (node, (_, kill))) =
wenzelm@22056
   421
      (kill (the_global_theory (History.current node)); toplevel)
wenzelm@23363
   422
  | apply_tr _ _ (History f) (State (node, term)) = State (f node, term)
wenzelm@23363
   423
  | apply_tr int _ (Keep f) state = keep_state int f state
wenzelm@23363
   424
  | apply_tr int pos (Transaction (hist, f)) (State state) =
wenzelm@23363
   425
      transaction hist pos (fn x => f int x) state
wenzelm@23363
   426
  | apply_tr _ _ _ _ = raise UNDEF;
wenzelm@6689
   427
wenzelm@23363
   428
fun apply_union _ _ [] state = raise FAILURE (state, UNDEF)
wenzelm@23363
   429
  | apply_union int pos (tr :: trs) state =
wenzelm@23363
   430
      apply_tr int pos tr state
wenzelm@23363
   431
        handle UNDEF => apply_union int pos trs state
wenzelm@23363
   432
          | FAILURE (alt_state, UNDEF) => apply_union int pos trs alt_state
wenzelm@6689
   433
          | exn as FAILURE _ => raise exn
wenzelm@6689
   434
          | exn => raise FAILURE (state, exn);
wenzelm@6689
   435
wenzelm@6689
   436
in
wenzelm@6689
   437
wenzelm@23363
   438
fun apply_trans int pos trs state = (apply_union int pos trs state, NONE)
skalberg@15531
   439
  handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn);
wenzelm@6689
   440
wenzelm@6689
   441
end;
wenzelm@5828
   442
wenzelm@5828
   443
wenzelm@5828
   444
(* datatype transition *)
wenzelm@5828
   445
wenzelm@5828
   446
datatype transition = Transition of
wenzelm@16815
   447
 {name: string,                        (*command name*)
wenzelm@16815
   448
  pos: Position.T,                     (*source position*)
wenzelm@16815
   449
  source: OuterLex.token list option,  (*source text*)
wenzelm@16815
   450
  int_only: bool,                      (*interactive-only*)
wenzelm@16815
   451
  print: string list,                  (*print modes (union)*)
wenzelm@16815
   452
  no_timing: bool,                     (*suppress timing*)
wenzelm@16815
   453
  trans: trans list};                  (*primitive transitions (union)*)
wenzelm@5828
   454
wenzelm@14923
   455
fun make_transition (name, pos, source, int_only, print, no_timing, trans) =
wenzelm@14923
   456
  Transition {name = name, pos = pos, source = source,
wenzelm@14923
   457
    int_only = int_only, print = print, no_timing = no_timing, trans = trans};
wenzelm@5828
   458
wenzelm@14923
   459
fun map_transition f (Transition {name, pos, source, int_only, print, no_timing, trans}) =
wenzelm@14923
   460
  make_transition (f (name, pos, source, int_only, print, no_timing, trans));
wenzelm@5828
   461
wenzelm@16607
   462
val empty = make_transition ("<unknown>", Position.none, NONE, false, [], false, []);
wenzelm@14923
   463
wenzelm@14923
   464
fun name_of (Transition {name, ...}) = name;
wenzelm@14923
   465
fun source_of (Transition {source, ...}) = source;
wenzelm@5828
   466
wenzelm@5828
   467
wenzelm@5828
   468
(* diagnostics *)
wenzelm@5828
   469
wenzelm@5828
   470
fun str_of_transition (Transition {name, pos, ...}) = quote name ^ Position.str_of pos;
wenzelm@5828
   471
wenzelm@5828
   472
fun command_msg msg tr = msg ^ "command " ^ str_of_transition tr;
wenzelm@5828
   473
fun at_command tr = command_msg "At " tr ^ ".";
wenzelm@5828
   474
wenzelm@5828
   475
fun type_error tr state =
wenzelm@18685
   476
  ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state);
wenzelm@5828
   477
wenzelm@5828
   478
wenzelm@5828
   479
(* modify transitions *)
wenzelm@5828
   480
wenzelm@14923
   481
fun name nm = map_transition (fn (_, pos, source, int_only, print, no_timing, trans) =>
wenzelm@14923
   482
  (nm, pos, source, int_only, print, no_timing, trans));
wenzelm@5828
   483
wenzelm@14923
   484
fun position pos = map_transition (fn (name, _, source, int_only, print, no_timing, trans) =>
wenzelm@14923
   485
  (name, pos, source, int_only, print, no_timing, trans));
wenzelm@5828
   486
wenzelm@14923
   487
fun source src = map_transition (fn (name, pos, _, int_only, print, no_timing, trans) =>
skalberg@15531
   488
  (name, pos, SOME src, int_only, print, no_timing, trans));
wenzelm@5828
   489
wenzelm@14923
   490
fun interactive int_only = map_transition (fn (name, pos, source, _, print, no_timing, trans) =>
wenzelm@14923
   491
  (name, pos, source, int_only, print, no_timing, trans));
wenzelm@5828
   492
wenzelm@17363
   493
val no_timing = map_transition (fn (name, pos, source, int_only, print, _, trans) =>
wenzelm@17363
   494
  (name, pos, source, int_only, print, true, trans));
wenzelm@17363
   495
wenzelm@17363
   496
fun add_trans tr = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) =>
wenzelm@17363
   497
  (name, pos, source, int_only, print, no_timing, trans @ [tr]));
wenzelm@17363
   498
wenzelm@16607
   499
fun print' mode = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) =>
wenzelm@16607
   500
  (name, pos, source, int_only, insert (op =) mode print, no_timing, trans));
wenzelm@16607
   501
wenzelm@16607
   502
val print = print' "";
wenzelm@9010
   503
wenzelm@17363
   504
val three_buffersN = "three_buffers";
wenzelm@17363
   505
val print3 = print' three_buffersN;
wenzelm@5828
   506
wenzelm@5828
   507
wenzelm@21007
   508
(* basic transitions *)
wenzelm@5828
   509
wenzelm@22056
   510
fun init_theory f exit kill = add_trans (Init (f, (exit, kill)));
wenzelm@21958
   511
val init_empty = add_trans o InitEmpty;
wenzelm@6689
   512
val exit = add_trans Exit;
wenzelm@21958
   513
val undo_exit = add_trans UndoExit;
wenzelm@6689
   514
val kill = add_trans Kill;
wenzelm@20128
   515
val history = add_trans o History;
wenzelm@7612
   516
val keep' = add_trans o Keep;
wenzelm@18592
   517
fun map_current f = add_trans (Transaction (false, f));
wenzelm@18592
   518
fun app_current f = add_trans (Transaction (true, f));
wenzelm@5828
   519
wenzelm@7612
   520
fun keep f = add_trans (Keep (fn _ => f));
wenzelm@5828
   521
fun imperative f = keep (fn _ => f ());
wenzelm@5828
   522
wenzelm@21007
   523
val unknown_theory = imperative (fn () => warning "Unknown theory context");
wenzelm@21007
   524
val unknown_proof = imperative (fn () => warning "Unknown proof context");
wenzelm@21007
   525
val unknown_context = imperative (fn () => warning "Unknown context");
wenzelm@15668
   526
wenzelm@21007
   527
wenzelm@21007
   528
(* theory transitions *)
wenzelm@15668
   529
wenzelm@20963
   530
fun theory' f = app_current (fn int =>
wenzelm@20963
   531
  (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (f int thy), NONE)
wenzelm@20963
   532
    | _ => raise UNDEF));
wenzelm@18563
   533
wenzelm@20963
   534
fun theory f = theory' (K f);
wenzelm@17076
   535
wenzelm@21294
   536
fun begin_local_theory begin f = app_current (fn _ =>
wenzelm@20963
   537
  (fn Theory (Context.Theory thy, _) =>
wenzelm@20963
   538
        let
wenzelm@20985
   539
          val lthy = f thy;
wenzelm@21294
   540
          val gthy = if begin then Context.Proof lthy else Context.Theory (loc_exit lthy);
wenzelm@21294
   541
        in Theory (gthy, SOME lthy) end
wenzelm@20963
   542
    | _ => raise UNDEF));
wenzelm@20963
   543
wenzelm@21294
   544
val end_local_theory = app_current (fn _ =>
wenzelm@21294
   545
  (fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (loc_exit lthy), SOME lthy)
wenzelm@21007
   546
    | _ => raise UNDEF));
wenzelm@21007
   547
wenzelm@21007
   548
local
wenzelm@21007
   549
wenzelm@20963
   550
fun local_theory_presentation loc f g = app_current (fn int =>
wenzelm@21294
   551
  (fn Theory (gthy, _) =>
wenzelm@21294
   552
        let
wenzelm@24780
   553
          val pos = ContextPosition.get (Context.proof_of gthy);
wenzelm@21294
   554
          val finish = loc_finish loc gthy;
wenzelm@24795
   555
          val lthy' = f (ContextPosition.put_ctxt pos (loc_begin loc gthy));
wenzelm@21294
   556
        in Theory (finish lthy', SOME lthy') end
wenzelm@20963
   557
    | _ => raise UNDEF) #> tap (g int));
wenzelm@20963
   558
wenzelm@21007
   559
in
wenzelm@21007
   560
wenzelm@20963
   561
fun local_theory loc f = local_theory_presentation loc f (K I);
wenzelm@20963
   562
fun present_local_theory loc g = local_theory_presentation loc I g;
wenzelm@18955
   563
wenzelm@21007
   564
end;
wenzelm@21007
   565
wenzelm@21007
   566
wenzelm@21007
   567
(* proof transitions *)
wenzelm@21007
   568
wenzelm@21007
   569
fun end_proof f = map_current (fn int =>
wenzelm@24795
   570
  (fn Proof (prf, (finish, _)) =>
wenzelm@21007
   571
        let val state = ProofHistory.current prf in
wenzelm@21007
   572
          if can (Proof.assert_bottom true) state then
wenzelm@21007
   573
            let
wenzelm@21007
   574
              val ctxt' = f int state;
wenzelm@21007
   575
              val gthy' = finish ctxt';
wenzelm@21007
   576
            in Theory (gthy', SOME ctxt') end
wenzelm@21007
   577
          else raise UNDEF
wenzelm@21007
   578
        end
wenzelm@21007
   579
    | SkipProof (h, (gthy, _)) =>
wenzelm@21007
   580
        if History.current h = 0 then Theory (gthy, NONE) else raise UNDEF
wenzelm@21007
   581
    | _ => raise UNDEF));
wenzelm@21007
   582
wenzelm@21294
   583
local
wenzelm@21294
   584
wenzelm@21294
   585
fun begin_proof init finish = app_current (fn int =>
wenzelm@21294
   586
  (fn Theory (gthy, _) =>
wenzelm@21294
   587
    let
berghofe@24453
   588
      val prf = init int gthy;
wenzelm@21294
   589
      val schematic = Proof.schematic_goal prf;
wenzelm@21294
   590
    in
wenzelm@21294
   591
      if ! skip_proofs andalso schematic then
wenzelm@21294
   592
        warning "Cannot skip proof of schematic goal statement"
wenzelm@21294
   593
      else ();
wenzelm@21294
   594
      if ! skip_proofs andalso not schematic then
wenzelm@21294
   595
        SkipProof
wenzelm@21294
   596
          (History.init (undo_limit int) 0, (finish gthy (Proof.global_skip_proof int prf), gthy))
wenzelm@21294
   597
      else Proof (ProofHistory.init (undo_limit int) prf, (finish gthy, gthy))
wenzelm@21294
   598
    end
wenzelm@21294
   599
  | _ => raise UNDEF));
wenzelm@21294
   600
wenzelm@21294
   601
in
wenzelm@21294
   602
wenzelm@24780
   603
fun local_theory_to_proof' loc f = begin_proof
wenzelm@24780
   604
  (fn int => fn gthy =>
wenzelm@24795
   605
    f int (ContextPosition.put_ctxt (ContextPosition.get (Context.proof_of gthy))
wenzelm@24795
   606
      (loc_begin loc gthy)))
wenzelm@24780
   607
  (loc_finish loc);
wenzelm@24780
   608
berghofe@24453
   609
fun local_theory_to_proof loc f = local_theory_to_proof' loc (K f);
wenzelm@21294
   610
wenzelm@21294
   611
fun theory_to_proof f = begin_proof
wenzelm@24780
   612
  (K (fn Context.Theory thy => f thy | _ => raise UNDEF))
wenzelm@24780
   613
  (K (Context.Theory o ProofContext.theory_of));
wenzelm@21294
   614
wenzelm@21294
   615
end;
wenzelm@21294
   616
wenzelm@21007
   617
val forget_proof = map_current (fn _ =>
wenzelm@21007
   618
  (fn Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
wenzelm@21007
   619
    | SkipProof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
wenzelm@21007
   620
    | _ => raise UNDEF));
wenzelm@21007
   621
wenzelm@21177
   622
fun present_proof f = map_current (fn int =>
wenzelm@21177
   623
  (fn Proof (prf, x) => Proof (ProofHistory.apply I prf, x)
wenzelm@21177
   624
    | SkipProof (h, x) => SkipProof (History.apply I h, x)
wenzelm@21177
   625
    | _ => raise UNDEF) #> tap (f int));
wenzelm@21177
   626
wenzelm@17904
   627
fun proofs' f = map_current (fn int =>
wenzelm@21007
   628
  (fn Proof (prf, x) => Proof (ProofHistory.applys (f int) prf, x)
wenzelm@21007
   629
    | SkipProof (h, x) => SkipProof (History.apply I h, x)
wenzelm@16815
   630
    | _ => raise UNDEF));
wenzelm@15668
   631
wenzelm@17904
   632
fun proof' f = proofs' (Seq.single oo f);
wenzelm@17904
   633
val proofs = proofs' o K;
wenzelm@6689
   634
val proof = proof' o K;
wenzelm@16815
   635
wenzelm@16815
   636
fun actual_proof f = map_current (fn _ =>
wenzelm@21007
   637
  (fn Proof (prf, x) => Proof (f prf, x)
wenzelm@20963
   638
    | _ => raise UNDEF));
wenzelm@16815
   639
wenzelm@16815
   640
fun skip_proof f = map_current (fn _ =>
wenzelm@21007
   641
  (fn SkipProof (h, x) => SkipProof (f h, x)
wenzelm@18563
   642
    | _ => raise UNDEF));
wenzelm@18563
   643
wenzelm@16815
   644
fun skip_proof_to_theory p = map_current (fn _ =>
wenzelm@21007
   645
  (fn SkipProof (h, (gthy, _)) =>
wenzelm@21007
   646
    if p (History.current h) then Theory (gthy, NONE)
wenzelm@17904
   647
    else raise UNDEF
wenzelm@17904
   648
  | _ => raise UNDEF));
wenzelm@5828
   649
wenzelm@5828
   650
wenzelm@5828
   651
wenzelm@5828
   652
(** toplevel transactions **)
wenzelm@5828
   653
wenzelm@5828
   654
(* apply transitions *)
wenzelm@5828
   655
wenzelm@6664
   656
local
wenzelm@6664
   657
wenzelm@23363
   658
fun app int (tr as Transition {trans, pos, int_only, print, no_timing, ...}) state =
wenzelm@5828
   659
  let
wenzelm@21962
   660
    val _ =
wenzelm@21962
   661
      if not int andalso int_only then warning (command_msg "Interactive-only " tr)
wenzelm@21962
   662
      else ();
wenzelm@16682
   663
wenzelm@22588
   664
    fun do_timing f x = (warning (command_msg "" tr); timeap f x);
wenzelm@16682
   665
    fun do_profiling f x = profile (! profiling) f x;
wenzelm@16682
   666
wenzelm@6689
   667
    val (result, opt_exn) =
wenzelm@23363
   668
       state |> (apply_trans int pos trans
wenzelm@23428
   669
        |> (if ! profiling > 0 andalso not no_timing then do_profiling else I)
wenzelm@16729
   670
        |> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I));
wenzelm@21962
   671
    val _ =
wenzelm@24634
   672
      if int andalso not (! quiet) andalso exists (member (op =) print) ("" :: print_mode_value ())
wenzelm@21962
   673
      then print_state false result else ();
skalberg@15570
   674
  in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) opt_exn) end;
wenzelm@6664
   675
wenzelm@6664
   676
in
wenzelm@5828
   677
wenzelm@6689
   678
fun apply int tr st =
wenzelm@6965
   679
  (case app int tr st of
skalberg@15531
   680
    (_, SOME TERMINATE) => NONE
skalberg@15531
   681
  | (_, SOME RESTART) => SOME (toplevel, NONE)
skalberg@15531
   682
  | (state', SOME (EXCURSION_FAIL exn_info)) => SOME (state', SOME exn_info)
skalberg@15531
   683
  | (state', SOME exn) => SOME (state', SOME (exn, at_command tr))
skalberg@15531
   684
  | (state', NONE) => SOME (state', NONE));
wenzelm@6664
   685
wenzelm@6664
   686
end;
wenzelm@5828
   687
wenzelm@5828
   688
wenzelm@17076
   689
(* excursion: toplevel -- apply transformers/presentation -- toplevel *)
wenzelm@5828
   690
wenzelm@6664
   691
local
wenzelm@6664
   692
wenzelm@5828
   693
fun excur [] x = x
wenzelm@17076
   694
  | excur ((tr, pr) :: trs) (st, res) =
wenzelm@17321
   695
      (case apply (! interact) tr st of
skalberg@15531
   696
        SOME (st', NONE) =>
wenzelm@18685
   697
          excur trs (st', pr st st' res handle exn =>
wenzelm@10324
   698
            raise EXCURSION_FAIL (exn, "Presentation failed\n" ^ at_command tr))
skalberg@15531
   699
      | SOME (st', SOME exn_info) => raise EXCURSION_FAIL exn_info
skalberg@15531
   700
      | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr));
wenzelm@5828
   701
wenzelm@17076
   702
fun no_pr _ _ _ = ();
wenzelm@17076
   703
wenzelm@6664
   704
in
wenzelm@6664
   705
wenzelm@17076
   706
fun present_excursion trs res =
wenzelm@21958
   707
  (case excur trs (toplevel, res) of
wenzelm@22056
   708
    (state as Toplevel _, res') => (safe_exit state; res')
wenzelm@18685
   709
  | _ => error "Unfinished development at end of input")
wenzelm@9134
   710
  handle exn => error (exn_message exn);
wenzelm@9134
   711
wenzelm@17076
   712
fun excursion trs = present_excursion (map (rpair no_pr) trs) ();
wenzelm@7062
   713
wenzelm@6664
   714
end;
wenzelm@6664
   715
wenzelm@5828
   716
wenzelm@5828
   717
wenzelm@5828
   718
(** interactive transformations **)
wenzelm@5828
   719
wenzelm@5828
   720
(* the global state reference *)
wenzelm@5828
   721
skalberg@15531
   722
val global_state = ref (toplevel, NONE: (exn * string) option);
wenzelm@5828
   723
skalberg@15531
   724
fun set_state state = global_state := (state, NONE);
wenzelm@5828
   725
fun get_state () = fst (! global_state);
wenzelm@5828
   726
fun exn () = snd (! global_state);
wenzelm@5828
   727
wenzelm@5828
   728
wenzelm@24295
   729
(* apply transformers to global state --- NOT THREAD-SAFE! *)
wenzelm@5828
   730
wenzelm@14985
   731
nonfix >> >>>;
wenzelm@5828
   732
wenzelm@24295
   733
fun >> tr =
wenzelm@5828
   734
  (case apply true tr (get_state ()) of
skalberg@15531
   735
    NONE => false
skalberg@15531
   736
  | SOME (state', exn_info) =>
wenzelm@5828
   737
      (global_state := (state', exn_info);
wenzelm@22014
   738
        print_exn exn_info;
wenzelm@24295
   739
        true));
wenzelm@5828
   740
wenzelm@14985
   741
fun >>> [] = ()
wenzelm@14985
   742
  | >>> (tr :: trs) = if >> tr then >>> trs else ();
wenzelm@14985
   743
wenzelm@21958
   744
fun init_state () = (>> (init_empty (K ()) empty); ());
wenzelm@21958
   745
wenzelm@21958
   746
wenzelm@21958
   747
(* the Isar source of transitions *)
wenzelm@21958
   748
wenzelm@21958
   749
type 'a isar =
wenzelm@21958
   750
  (transition, (transition option,
wenzelm@21958
   751
    (OuterLex.token, (OuterLex.token option, (OuterLex.token, (OuterLex.token,
wenzelm@21958
   752
      Position.T * (Symbol.symbol, (string, 'a) Source.source) Source.source)
wenzelm@21958
   753
          Source.source) Source.source) Source.source) Source.source) Source.source) Source.source;
wenzelm@21958
   754
wenzelm@17513
   755
(*Spurious interrupts ahead!  Race condition?*)
wenzelm@17513
   756
fun get_interrupt src = SOME (Source.get_single src) handle Interrupt => NONE;
wenzelm@7602
   757
wenzelm@20928
   758
fun warn_secure () =
wenzelm@20928
   759
  let val secure = Secure.is_secure ()
wenzelm@20928
   760
  in if secure then warning "Cannot exit to ML in secure mode" else (); secure end;
wenzelm@20928
   761
wenzelm@5828
   762
fun raw_loop src =
wenzelm@23720
   763
  let val prompt = Output.escape (Markup.enclose Markup.prompt Source.default_prompt) in
wenzelm@23640
   764
    (case get_interrupt (Source.set_prompt prompt src) of
wenzelm@23640
   765
      NONE => (writeln "\nInterrupt."; raw_loop src)
wenzelm@23640
   766
    | SOME NONE => if warn_secure () then quit () else ()
wenzelm@23640
   767
    | SOME (SOME (tr, src')) =>
wenzelm@23640
   768
        if >> tr orelse warn_secure () then raw_loop src'
wenzelm@23640
   769
        else ())
wenzelm@23640
   770
  end;
wenzelm@5828
   771
wenzelm@12987
   772
fun loop src = ignore_interrupt raw_loop src;
wenzelm@5828
   773
wenzelm@5828
   774
end;