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