src/Pure/Isar/toplevel.ML
author wenzelm
Wed, 01 Oct 2008 12:00:02 +0200
changeset 28443 de653f1ad78b
parent 28433 b3dab95f098f
child 28446 a01de3b3fa2e
permissions -rw-r--r--
more robust treatment of Interrupt (cf. exn.ML);
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@26602
     5
Isabelle/Isar toplevel transactions.
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@27564
    14
  val proof_node: node -> ProofNode.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@26602
    18
  val toplevel: state
wenzelm@7732
    19
  val is_toplevel: state -> bool
wenzelm@18589
    20
  val is_theory: state -> bool
wenzelm@18589
    21
  val is_proof: state -> bool
wenzelm@17076
    22
  val level: state -> int
wenzelm@27576
    23
  val previous_node_of: state -> node option
wenzelm@5828
    24
  val node_of: state -> node
wenzelm@20963
    25
  val node_case: (generic_theory -> 'a) -> (Proof.state -> 'a) -> state -> 'a
wenzelm@21506
    26
  val context_of: state -> Proof.context
wenzelm@22089
    27
  val generic_theory_of: state -> generic_theory
wenzelm@5828
    28
  val theory_of: state -> theory
wenzelm@5828
    29
  val proof_of: state -> Proof.state
wenzelm@18589
    30
  val proof_position_of: state -> int
wenzelm@21007
    31
  val enter_proof_body: state -> Proof.state
wenzelm@16815
    32
  val print_state_context: state -> unit
wenzelm@16815
    33
  val print_state: bool -> state -> unit
wenzelm@16682
    34
  val quiet: bool ref
wenzelm@16682
    35
  val debug: bool ref
wenzelm@17321
    36
  val interact: bool ref
wenzelm@16682
    37
  val timing: bool ref
wenzelm@16682
    38
  val profiling: int ref
wenzelm@16815
    39
  val skip_proofs: bool ref
wenzelm@5828
    40
  exception TERMINATE
wenzelm@27583
    41
  exception TOPLEVEL_ERROR
wenzelm@26256
    42
  exception CONTEXT of Proof.context * exn
wenzelm@20128
    43
  val exn_message: exn -> string
wenzelm@20128
    44
  val program: (unit -> 'a) -> 'a
wenzelm@16682
    45
  type transition
wenzelm@5828
    46
  val empty: transition
wenzelm@27441
    47
  val init_of: transition -> string option
wenzelm@27427
    48
  val name_of: transition -> string
wenzelm@28105
    49
  val pos_of: transition -> Position.T
wenzelm@27500
    50
  val str_of: transition -> string
wenzelm@5828
    51
  val name: string -> transition -> transition
wenzelm@5828
    52
  val position: Position.T -> transition -> transition
wenzelm@5828
    53
  val interactive: bool -> transition -> transition
wenzelm@5828
    54
  val print: transition -> transition
wenzelm@9010
    55
  val no_timing: transition -> transition
wenzelm@27576
    56
  val init_theory: string -> (bool -> theory) -> (theory -> unit) -> transition -> transition
wenzelm@6689
    57
  val exit: transition -> transition
wenzelm@5828
    58
  val keep: (state -> unit) -> transition -> transition
wenzelm@7612
    59
  val keep': (bool -> state -> unit) -> transition -> transition
wenzelm@5828
    60
  val imperative: (unit -> unit) -> transition -> transition
wenzelm@27840
    61
  val ignored: Position.T -> transition
wenzelm@27840
    62
  val malformed: Position.T -> string -> transition
wenzelm@5828
    63
  val theory: (theory -> theory) -> transition -> transition
wenzelm@26491
    64
  val generic_theory: (generic_theory -> generic_theory) -> transition -> transition
wenzelm@7612
    65
  val theory': (bool -> theory -> theory) -> transition -> transition
wenzelm@20985
    66
  val begin_local_theory: bool -> (theory -> local_theory) -> transition -> transition
wenzelm@21007
    67
  val end_local_theory: transition -> transition
wenzelm@20963
    68
  val local_theory: xstring option -> (local_theory -> local_theory) -> transition -> transition
wenzelm@27859
    69
  val present_local_theory: xstring option -> (node -> unit) -> transition -> transition
berghofe@24453
    70
  val local_theory_to_proof': xstring option -> (bool -> local_theory -> Proof.state) ->
berghofe@24453
    71
    transition -> transition
wenzelm@21007
    72
  val local_theory_to_proof: xstring option -> (local_theory -> Proof.state) ->
wenzelm@21007
    73
    transition -> transition
wenzelm@17363
    74
  val theory_to_proof: (theory -> Proof.state) -> transition -> transition
wenzelm@21007
    75
  val end_proof: (bool -> Proof.state -> Proof.context) -> transition -> transition
wenzelm@21007
    76
  val forget_proof: transition -> transition
wenzelm@27859
    77
  val present_proof: (node -> unit) -> transition -> transition
wenzelm@21177
    78
  val proofs': (bool -> Proof.state -> Proof.state Seq.seq) -> transition -> transition
wenzelm@21177
    79
  val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition
wenzelm@21177
    80
  val proofs: (Proof.state -> Proof.state Seq.seq) -> transition -> transition
wenzelm@17904
    81
  val proof: (Proof.state -> Proof.state) -> transition -> transition
wenzelm@27564
    82
  val actual_proof: (ProofNode.T -> ProofNode.T) -> transition -> transition
wenzelm@27564
    83
  val skip_proof: (int -> int) -> transition -> transition
wenzelm@17904
    84
  val skip_proof_to_theory: (int -> bool) -> transition -> transition
wenzelm@27427
    85
  val get_id: transition -> string option
wenzelm@27427
    86
  val put_id: string -> transition -> transition
wenzelm@9512
    87
  val unknown_theory: transition -> transition
wenzelm@9512
    88
  val unknown_proof: transition -> transition
wenzelm@9512
    89
  val unknown_context: transition -> transition
wenzelm@28425
    90
  val setmp_thread_position: transition -> ('a -> 'b) -> 'a -> 'b
wenzelm@27606
    91
  val status: transition -> Markup.T -> unit
wenzelm@26602
    92
  val error_msg: transition -> exn * string -> unit
wenzelm@28103
    93
  val add_hook: (transition -> state -> state -> unit) -> unit
wenzelm@26602
    94
  val transition: bool -> transition -> state -> (state * (exn * string) option) option
wenzelm@28425
    95
  val commit_exit: Position.T -> transition
wenzelm@28425
    96
  val command: transition -> state -> state
wenzelm@28433
    97
  val excursion: (transition * transition list) list -> (transition * state) list * (unit -> unit)
wenzelm@5828
    98
end;
wenzelm@5828
    99
wenzelm@6965
   100
structure Toplevel: TOPLEVEL =
wenzelm@5828
   101
struct
wenzelm@5828
   102
wenzelm@5828
   103
(** toplevel state **)
wenzelm@5828
   104
wenzelm@19063
   105
exception UNDEF;
wenzelm@19063
   106
wenzelm@19063
   107
wenzelm@21294
   108
(* local theory wrappers *)
wenzelm@5828
   109
wenzelm@20963
   110
type generic_theory = Context.generic;    (*theory or local_theory*)
wenzelm@20963
   111
wenzelm@25292
   112
val loc_init = TheoryTarget.context;
wenzelm@28394
   113
val loc_exit = LocalTheory.exit_global;
wenzelm@21294
   114
wenzelm@25292
   115
fun loc_begin loc (Context.Theory thy) = loc_init (the_default "-" loc) thy
wenzelm@21294
   116
  | loc_begin NONE (Context.Proof lthy) = lthy
haftmann@25269
   117
  | loc_begin (SOME loc) (Context.Proof lthy) = loc_init loc (loc_exit lthy);
wenzelm@21294
   118
wenzelm@21294
   119
fun loc_finish _ (Context.Theory _) = Context.Theory o loc_exit
wenzelm@21294
   120
  | loc_finish NONE (Context.Proof _) = Context.Proof o LocalTheory.restore
wenzelm@25292
   121
  | loc_finish (SOME _) (Context.Proof lthy) = fn lthy' =>
wenzelm@25292
   122
      Context.Proof (LocalTheory.reinit (LocalTheory.raw_theory (K (loc_exit lthy')) lthy));
wenzelm@21294
   123
wenzelm@21294
   124
wenzelm@21958
   125
(* datatype node *)
wenzelm@21294
   126
wenzelm@5828
   127
datatype node =
wenzelm@27576
   128
  Theory of generic_theory * Proof.context option
wenzelm@27576
   129
    (*theory with presentation context*) |
wenzelm@27576
   130
  Proof of ProofNode.T * ((Proof.context -> generic_theory) * generic_theory)
wenzelm@27576
   131
    (*proof node, finish, original theory*) |
wenzelm@27564
   132
  SkipProof of int * (generic_theory * generic_theory);
wenzelm@27564
   133
    (*proof depth, resulting theory, original theory*)
wenzelm@5828
   134
wenzelm@22056
   135
val the_global_theory = fn Theory (Context.Theory thy, _) => thy | _ => raise UNDEF;
wenzelm@20963
   136
val theory_node = fn Theory (gthy, _) => SOME gthy | _ => NONE;
wenzelm@18589
   137
val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE;
wenzelm@18589
   138
wenzelm@20963
   139
fun cases_node f _ (Theory (gthy, _)) = f gthy
wenzelm@27564
   140
  | cases_node _ g (Proof (prf, _)) = g (ProofNode.current prf)
wenzelm@21007
   141
  | cases_node f _ (SkipProof (_, (gthy, _))) = f gthy;
wenzelm@19063
   142
wenzelm@20963
   143
fun presentation_context (SOME (Theory (_, SOME ctxt))) NONE = ctxt
wenzelm@20963
   144
  | presentation_context (SOME node) NONE = cases_node Context.proof_of Proof.context_of node
wenzelm@20963
   145
  | presentation_context (SOME node) (SOME loc) =
haftmann@25269
   146
      loc_init loc (cases_node Context.theory_of Proof.theory_of node)
wenzelm@20963
   147
  | presentation_context NONE _ = raise UNDEF;
wenzelm@19063
   148
wenzelm@26624
   149
fun reset_presentation (Theory (gthy, _)) = Theory (gthy, NONE)
wenzelm@26624
   150
  | reset_presentation node = node;
wenzelm@26624
   151
wenzelm@5828
   152
wenzelm@21958
   153
(* datatype state *)
wenzelm@5828
   154
wenzelm@27601
   155
type node_info = node * (theory -> unit);  (*node with exit operation*)
wenzelm@27601
   156
datatype state = State of node_info option * node_info option;  (*current, previous*)
wenzelm@21958
   157
wenzelm@27576
   158
val toplevel = State (NONE, NONE);
wenzelm@21958
   159
wenzelm@27576
   160
fun is_toplevel (State (NONE, _)) = true
wenzelm@7732
   161
  | is_toplevel _ = false;
wenzelm@7732
   162
wenzelm@27576
   163
fun level (State (NONE, _)) = 0
wenzelm@27576
   164
  | level (State (SOME (Theory _, _), _)) = 0
wenzelm@27576
   165
  | level (State (SOME (Proof (prf, _), _), _)) = Proof.level (ProofNode.current prf)
wenzelm@27576
   166
  | level (State (SOME (SkipProof (d, _), _), _)) = d + 1;   (*different notion of proof depth!*)
wenzelm@17076
   167
wenzelm@27576
   168
fun str_of_state (State (NONE, _)) = "at top level"
wenzelm@27576
   169
  | str_of_state (State (SOME (Theory (Context.Theory _, _), _), _)) = "in theory mode"
wenzelm@27576
   170
  | str_of_state (State (SOME (Theory (Context.Proof _, _), _), _)) = "in local theory mode"
wenzelm@27576
   171
  | str_of_state (State (SOME (Proof _, _), _)) = "in proof mode"
wenzelm@27576
   172
  | str_of_state (State (SOME (SkipProof _, _), _)) = "in skipped proof mode";
wenzelm@5946
   173
wenzelm@5946
   174
wenzelm@27576
   175
(* current node *)
wenzelm@5828
   176
wenzelm@27603
   177
fun previous_node_of (State (_, prev)) = Option.map #1 prev;
wenzelm@6689
   178
wenzelm@27576
   179
fun node_of (State (NONE, _)) = raise UNDEF
wenzelm@27576
   180
  | node_of (State (SOME (node, _), _)) = node;
wenzelm@5828
   181
wenzelm@18589
   182
fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state));
wenzelm@18589
   183
fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state));
wenzelm@18589
   184
wenzelm@19063
   185
fun node_case f g state = cases_node f g (node_of state);
wenzelm@5828
   186
wenzelm@21506
   187
val context_of = node_case Context.proof_of Proof.context_of;
wenzelm@22089
   188
val generic_theory_of = node_case I (Context.Proof o Proof.context_of);
wenzelm@20963
   189
val theory_of = node_case Context.theory_of Proof.theory_of;
wenzelm@18589
   190
val proof_of = node_case (fn _ => raise UNDEF) I;
wenzelm@17208
   191
wenzelm@18589
   192
fun proof_position_of state =
wenzelm@18589
   193
  (case node_of state of
wenzelm@27564
   194
    Proof (prf, _) => ProofNode.position prf
wenzelm@18589
   195
  | _ => raise UNDEF);
wenzelm@6664
   196
wenzelm@21007
   197
val enter_proof_body = node_case (Proof.init o Context.proof_of) Proof.enter_forward;
wenzelm@5828
   198
wenzelm@5828
   199
wenzelm@16815
   200
(* print state *)
wenzelm@16815
   201
wenzelm@25292
   202
val pretty_context = LocalTheory.pretty o Context.cases (TheoryTarget.init NONE) I;
wenzelm@16815
   203
wenzelm@23640
   204
fun print_state_context state =
wenzelm@24795
   205
  (case try node_of state of
wenzelm@21506
   206
    NONE => []
wenzelm@24795
   207
  | SOME (Theory (gthy, _)) => pretty_context gthy
wenzelm@24795
   208
  | SOME (Proof (_, (_, gthy))) => pretty_context gthy
wenzelm@24795
   209
  | SOME (SkipProof (_, (gthy, _))) => pretty_context gthy)
wenzelm@23640
   210
  |> Pretty.chunks |> Pretty.writeln;
wenzelm@16815
   211
wenzelm@23640
   212
fun print_state prf_only state =
wenzelm@23701
   213
  (case try node_of state of
wenzelm@23701
   214
    NONE => []
wenzelm@23701
   215
  | SOME (Theory (gthy, _)) => if prf_only then [] else pretty_context gthy
wenzelm@23701
   216
  | SOME (Proof (prf, _)) =>
wenzelm@27564
   217
      Proof.pretty_state (ProofNode.position prf) (ProofNode.current prf)
wenzelm@27564
   218
  | SOME (SkipProof (d, _)) => [Pretty.str ("skipped proof: depth " ^ string_of_int d)])
wenzelm@23701
   219
  |> Pretty.markup_chunks Markup.state |> Pretty.writeln;
wenzelm@16815
   220
wenzelm@16815
   221
wenzelm@15668
   222
wenzelm@5828
   223
(** toplevel transitions **)
wenzelm@5828
   224
wenzelm@16682
   225
val quiet = ref false;
wenzelm@22135
   226
val debug = Output.debugging;
wenzelm@17321
   227
val interact = ref false;
wenzelm@16682
   228
val timing = Output.timing;
wenzelm@16682
   229
val profiling = ref 0;
wenzelm@16815
   230
val skip_proofs = ref false;
wenzelm@16682
   231
wenzelm@5828
   232
exception TERMINATE;
wenzelm@7022
   233
exception EXCURSION_FAIL of exn * string;
wenzelm@6689
   234
exception FAILURE of state * exn;
wenzelm@24055
   235
exception TOPLEVEL_ERROR;
wenzelm@5828
   236
wenzelm@20128
   237
wenzelm@20128
   238
(* print exceptions *)
wenzelm@20128
   239
wenzelm@26256
   240
exception CONTEXT of Proof.context * exn;
wenzelm@26256
   241
wenzelm@26256
   242
fun exn_context NONE exn = exn
wenzelm@26256
   243
  | exn_context (SOME ctxt) exn = CONTEXT (ctxt, exn);
wenzelm@26256
   244
wenzelm@20128
   245
local
wenzelm@20128
   246
wenzelm@26256
   247
fun if_context NONE _ _ = []
wenzelm@26256
   248
  | if_context (SOME ctxt) f xs = map (f ctxt) xs;
wenzelm@20128
   249
wenzelm@20128
   250
fun raised name [] = "exception " ^ name ^ " raised"
wenzelm@20128
   251
  | raised name [msg] = "exception " ^ name ^ " raised: " ^ msg
wenzelm@20128
   252
  | raised name msgs = cat_lines (("exception " ^ name ^ " raised:") :: msgs);
wenzelm@20128
   253
wenzelm@20128
   254
in
wenzelm@20128
   255
wenzelm@26256
   256
fun exn_message e =
wenzelm@26256
   257
  let
wenzelm@26256
   258
    val detailed = ! debug;
wenzelm@26256
   259
wenzelm@26256
   260
    fun exn_msg _ (CONTEXT (ctxt, exn)) = exn_msg (SOME ctxt) exn
wenzelm@26256
   261
      | exn_msg ctxt (Exn.EXCEPTIONS (exns, "")) = cat_lines (map (exn_msg ctxt) exns)
wenzelm@26256
   262
      | exn_msg ctxt (Exn.EXCEPTIONS (exns, msg)) = cat_lines (map (exn_msg ctxt) exns @ [msg])
wenzelm@26293
   263
      | exn_msg ctxt (EXCURSION_FAIL (exn, loc)) =
wenzelm@26293
   264
          exn_msg ctxt exn ^ Markup.markup Markup.location ("\n" ^ loc)
wenzelm@26256
   265
      | exn_msg _ TERMINATE = "Exit."
wenzelm@28443
   266
      | exn_msg _ Exn.Interrupt = "Interrupt."
wenzelm@26256
   267
      | exn_msg _ TimeLimit.TimeOut = "Timeout."
wenzelm@26256
   268
      | exn_msg _ TOPLEVEL_ERROR = "Error."
wenzelm@26256
   269
      | exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg
wenzelm@26256
   270
      | exn_msg _ (ERROR msg) = msg
wenzelm@26256
   271
      | exn_msg _ (Fail msg) = raised "Fail" [msg]
wenzelm@26256
   272
      | exn_msg _ (THEORY (msg, thys)) =
wenzelm@26256
   273
          raised "THEORY" (msg :: (if detailed then map Context.str_of_thy thys else []))
wenzelm@26256
   274
      | exn_msg _ (Syntax.AST (msg, asts)) = raised "AST" (msg ::
wenzelm@26256
   275
            (if detailed then map (Pretty.string_of o Syntax.pretty_ast) asts else []))
wenzelm@26256
   276
      | exn_msg ctxt (TYPE (msg, Ts, ts)) = raised "TYPE" (msg ::
wenzelm@26256
   277
            (if detailed then
wenzelm@26256
   278
              if_context ctxt Syntax.string_of_typ Ts @ if_context ctxt Syntax.string_of_term ts
wenzelm@26256
   279
             else []))
wenzelm@26256
   280
      | exn_msg ctxt (TERM (msg, ts)) = raised "TERM" (msg ::
wenzelm@26256
   281
            (if detailed then if_context ctxt Syntax.string_of_term ts else []))
wenzelm@26256
   282
      | exn_msg ctxt (THM (msg, i, thms)) = raised ("THM " ^ string_of_int i) (msg ::
wenzelm@26256
   283
            (if detailed then if_context ctxt ProofContext.string_of_thm thms else []))
wenzelm@26256
   284
      | exn_msg _ exn = raised (General.exnMessage exn) []
wenzelm@26256
   285
  in exn_msg NONE e end;
wenzelm@26256
   286
wenzelm@20128
   287
end;
wenzelm@20128
   288
wenzelm@20128
   289
wenzelm@20128
   290
(* controlled execution *)
wenzelm@20128
   291
wenzelm@20128
   292
local
wenzelm@20128
   293
wenzelm@18685
   294
fun debugging f x =
wenzelm@23940
   295
  if ! debug then exception_trace (fn () => f x)
wenzelm@18685
   296
  else f x;
wenzelm@18685
   297
wenzelm@26256
   298
fun toplevel_error f x =
wenzelm@26256
   299
  let val ctxt = try ML_Context.the_local_context () in
wenzelm@26256
   300
    f x handle exn =>
wenzelm@26256
   301
      (Output.error_msg (exn_message (exn_context ctxt exn)); raise TOPLEVEL_ERROR)
wenzelm@26256
   302
  end;
wenzelm@24055
   303
wenzelm@20128
   304
in
wenzelm@20128
   305
wenzelm@20128
   306
fun controlled_execution f =
wenzelm@20128
   307
  f
wenzelm@20128
   308
  |> debugging
wenzelm@24055
   309
  |> interruptible;
wenzelm@20128
   310
wenzelm@20128
   311
fun program f =
wenzelm@24055
   312
 (f
wenzelm@24055
   313
  |> debugging
wenzelm@24055
   314
  |> toplevel_error) ();
wenzelm@20128
   315
wenzelm@20128
   316
end;
wenzelm@20128
   317
wenzelm@5828
   318
wenzelm@27601
   319
(* node transactions -- maintaining stable checkpoints *)
wenzelm@7022
   320
wenzelm@6689
   321
local
wenzelm@6689
   322
wenzelm@26624
   323
fun is_draft_theory (Theory (gthy, _)) = Context.is_draft (Context.theory_of gthy)
wenzelm@26624
   324
  | is_draft_theory _ = false;
wenzelm@26624
   325
wenzelm@27601
   326
fun is_stale state = Context.is_stale (theory_of state) handle UNDEF => false;
wenzelm@27601
   327
wenzelm@26624
   328
fun stale_error NONE = SOME (ERROR "Stale theory encountered after successful execution!")
wenzelm@26624
   329
  | stale_error some = some;
wenzelm@16815
   330
wenzelm@27576
   331
fun map_theory f (Theory (gthy, ctxt)) =
wenzelm@27576
   332
      Theory (Context.mapping f (LocalTheory.raw_theory f) gthy, ctxt)
wenzelm@27576
   333
  | map_theory _ node = node;
wenzelm@6689
   334
wenzelm@6689
   335
in
wenzelm@6689
   336
wenzelm@27601
   337
fun apply_transaction pos f (node, exit) =
wenzelm@20128
   338
  let
wenzelm@27576
   339
    val _ = is_draft_theory node andalso error "Illegal draft theory in toplevel state";
wenzelm@27576
   340
    val cont_node = reset_presentation node;
wenzelm@27576
   341
    val back_node = map_theory (Theory.checkpoint o Theory.copy) cont_node;
wenzelm@27601
   342
    fun state_error e nd = (State (SOME (nd, exit), SOME (node, exit)), e);
wenzelm@6689
   343
wenzelm@20128
   344
    val (result, err) =
wenzelm@20128
   345
      cont_node
wenzelm@27576
   346
      |> controlled_execution f
wenzelm@26624
   347
      |> map_theory Theory.checkpoint
wenzelm@26624
   348
      |> state_error NONE
wenzelm@26624
   349
      handle exn => state_error (SOME exn) cont_node;
wenzelm@26624
   350
wenzelm@26624
   351
    val (result', err') =
wenzelm@26624
   352
      if is_stale result then state_error (stale_error err) back_node
wenzelm@26624
   353
      else (result, err);
wenzelm@20128
   354
  in
wenzelm@26624
   355
    (case err' of
wenzelm@26624
   356
      NONE => result'
wenzelm@26624
   357
    | SOME exn => raise FAILURE (result', exn))
wenzelm@20128
   358
  end;
wenzelm@6689
   359
wenzelm@6689
   360
end;
wenzelm@6689
   361
wenzelm@6689
   362
wenzelm@6689
   363
(* primitive transitions *)
wenzelm@6689
   364
wenzelm@5828
   365
datatype trans =
wenzelm@27576
   366
  Init of string * (bool -> theory) * (theory -> unit) | (*theory name, init, exit*)
wenzelm@27576
   367
  Exit |                                         (*formal exit of theory -- without committing*)
wenzelm@27576
   368
  CommitExit |                                   (*exit and commit final theory*)
wenzelm@27576
   369
  Keep of bool -> state -> unit |                (*peek at state*)
wenzelm@27576
   370
  Transaction of bool -> node -> node;           (*node transaction*)
wenzelm@21958
   371
wenzelm@6689
   372
local
wenzelm@6689
   373
wenzelm@27576
   374
fun apply_tr int _ (Init (_, f, exit)) (State (NONE, _)) =
wenzelm@27576
   375
      State (SOME (Theory (Context.Theory (Theory.checkpoint (f int)), NONE), exit), NONE)
wenzelm@27603
   376
  | apply_tr _ _ Exit (State (prev as SOME (Theory (Context.Theory _, _), _), _)) =
wenzelm@27603
   377
      State (NONE, prev)
wenzelm@27601
   378
  | apply_tr _ _ CommitExit (State (NONE, SOME (Theory (Context.Theory thy, _), exit))) =
wenzelm@27576
   379
      (controlled_execution exit thy; toplevel)
wenzelm@27576
   380
  | apply_tr int _ (Keep f) state =
wenzelm@27576
   381
      controlled_execution (fn x => tap (f int) x) state
wenzelm@27576
   382
  | apply_tr int pos (Transaction f) (State (SOME state, _)) =
wenzelm@27601
   383
      apply_transaction pos (fn x => f int x) state
wenzelm@23363
   384
  | apply_tr _ _ _ _ = raise UNDEF;
wenzelm@6689
   385
wenzelm@23363
   386
fun apply_union _ _ [] state = raise FAILURE (state, UNDEF)
wenzelm@23363
   387
  | apply_union int pos (tr :: trs) state =
wenzelm@23363
   388
      apply_tr int pos tr state
wenzelm@23363
   389
        handle UNDEF => apply_union int pos trs state
wenzelm@23363
   390
          | FAILURE (alt_state, UNDEF) => apply_union int pos trs alt_state
wenzelm@6689
   391
          | exn as FAILURE _ => raise exn
wenzelm@6689
   392
          | exn => raise FAILURE (state, exn);
wenzelm@6689
   393
wenzelm@6689
   394
in
wenzelm@6689
   395
wenzelm@23363
   396
fun apply_trans int pos trs state = (apply_union int pos trs state, NONE)
skalberg@15531
   397
  handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn);
wenzelm@6689
   398
wenzelm@6689
   399
end;
wenzelm@5828
   400
wenzelm@5828
   401
wenzelm@5828
   402
(* datatype transition *)
wenzelm@5828
   403
wenzelm@5828
   404
datatype transition = Transition of
wenzelm@26621
   405
 {name: string,              (*command name*)
wenzelm@26621
   406
  pos: Position.T,           (*source position*)
wenzelm@26621
   407
  int_only: bool,            (*interactive-only*)
wenzelm@26621
   408
  print: bool,               (*print result state*)
wenzelm@26621
   409
  no_timing: bool,           (*suppress timing*)
wenzelm@26621
   410
  trans: trans list};        (*primitive transitions (union)*)
wenzelm@5828
   411
wenzelm@26621
   412
fun make_transition (name, pos, int_only, print, no_timing, trans) =
wenzelm@26621
   413
  Transition {name = name, pos = pos, int_only = int_only, print = print, no_timing = no_timing,
wenzelm@26621
   414
    trans = trans};
wenzelm@5828
   415
wenzelm@26621
   416
fun map_transition f (Transition {name, pos, int_only, print, no_timing, trans}) =
wenzelm@26621
   417
  make_transition (f (name, pos, int_only, print, no_timing, trans));
wenzelm@5828
   418
wenzelm@27441
   419
val empty = make_transition ("", Position.none, false, false, false, []);
wenzelm@5828
   420
wenzelm@5828
   421
wenzelm@5828
   422
(* diagnostics *)
wenzelm@5828
   423
wenzelm@27441
   424
fun init_of (Transition {trans = [Init (name, _, _)], ...}) = SOME name
wenzelm@27441
   425
  | init_of _ = NONE;
wenzelm@27441
   426
wenzelm@27427
   427
fun name_of (Transition {name, ...}) = name;
wenzelm@28105
   428
fun pos_of (Transition {pos, ...}) = pos;
wenzelm@28105
   429
fun str_of tr = quote (name_of tr) ^ Position.str_of (pos_of tr);
wenzelm@5828
   430
wenzelm@27427
   431
fun command_msg msg tr = msg ^ "command " ^ str_of tr;
wenzelm@5828
   432
fun at_command tr = command_msg "At " tr ^ ".";
wenzelm@5828
   433
wenzelm@5828
   434
fun type_error tr state =
wenzelm@18685
   435
  ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state);
wenzelm@5828
   436
wenzelm@5828
   437
wenzelm@5828
   438
(* modify transitions *)
wenzelm@5828
   439
wenzelm@26621
   440
fun name nm = map_transition (fn (_, pos, int_only, print, no_timing, trans) =>
wenzelm@26621
   441
  (nm, pos, int_only, print, no_timing, trans));
wenzelm@5828
   442
wenzelm@26621
   443
fun position pos = map_transition (fn (name, _, int_only, print, no_timing, trans) =>
wenzelm@26621
   444
  (name, pos, int_only, print, no_timing, trans));
wenzelm@5828
   445
wenzelm@26621
   446
fun interactive int_only = map_transition (fn (name, pos, _, print, no_timing, trans) =>
wenzelm@26621
   447
  (name, pos, int_only, print, no_timing, trans));
wenzelm@5828
   448
wenzelm@26621
   449
val no_timing = map_transition (fn (name, pos, int_only, print, _, trans) =>
wenzelm@26621
   450
  (name, pos, int_only, print, true, trans));
wenzelm@5828
   451
wenzelm@26621
   452
fun add_trans tr = map_transition (fn (name, pos, int_only, print, no_timing, trans) =>
wenzelm@26621
   453
  (name, pos, int_only, print, no_timing, trans @ [tr]));
wenzelm@17363
   454
wenzelm@28433
   455
val reset_trans = map_transition (fn (name, pos, int_only, print, no_timing, _) =>
wenzelm@28433
   456
  (name, pos, int_only, print, no_timing, []));
wenzelm@28433
   457
wenzelm@26621
   458
val print = map_transition (fn (name, pos, int_only, _, no_timing, trans) =>
wenzelm@26621
   459
  (name, pos, int_only, true, no_timing, trans));
wenzelm@5828
   460
wenzelm@5828
   461
wenzelm@21007
   462
(* basic transitions *)
wenzelm@5828
   463
wenzelm@27576
   464
fun init_theory name f exit = add_trans (Init (name, f, exit));
wenzelm@6689
   465
val exit = add_trans Exit;
wenzelm@7612
   466
val keep' = add_trans o Keep;
wenzelm@27601
   467
fun transaction f = add_trans (Transaction f);
wenzelm@5828
   468
wenzelm@7612
   469
fun keep f = add_trans (Keep (fn _ => f));
wenzelm@5828
   470
fun imperative f = keep (fn _ => f ());
wenzelm@5828
   471
wenzelm@27840
   472
fun ignored pos = empty |> name "<ignored>" |> position pos |> imperative I;
wenzelm@27840
   473
fun malformed pos msg =
wenzelm@27840
   474
  empty |> name "<malformed>" |> position pos |> imperative (fn () => error msg);
wenzelm@27840
   475
wenzelm@21007
   476
val unknown_theory = imperative (fn () => warning "Unknown theory context");
wenzelm@21007
   477
val unknown_proof = imperative (fn () => warning "Unknown proof context");
wenzelm@21007
   478
val unknown_context = imperative (fn () => warning "Unknown context");
wenzelm@15668
   479
wenzelm@21007
   480
wenzelm@21007
   481
(* theory transitions *)
wenzelm@15668
   482
wenzelm@27601
   483
fun generic_theory f = transaction (fn _ =>
wenzelm@26491
   484
  (fn Theory (gthy, _) => Theory (f gthy, NONE)
wenzelm@26491
   485
    | _ => raise UNDEF));
wenzelm@26491
   486
wenzelm@27601
   487
fun theory' f = transaction (fn int =>
wenzelm@20963
   488
  (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (f int thy), NONE)
wenzelm@20963
   489
    | _ => raise UNDEF));
wenzelm@18563
   490
wenzelm@20963
   491
fun theory f = theory' (K f);
wenzelm@17076
   492
wenzelm@27601
   493
fun begin_local_theory begin f = transaction (fn _ =>
wenzelm@20963
   494
  (fn Theory (Context.Theory thy, _) =>
wenzelm@20963
   495
        let
wenzelm@20985
   496
          val lthy = f thy;
wenzelm@21294
   497
          val gthy = if begin then Context.Proof lthy else Context.Theory (loc_exit lthy);
wenzelm@21294
   498
        in Theory (gthy, SOME lthy) end
wenzelm@20963
   499
    | _ => raise UNDEF));
wenzelm@20963
   500
wenzelm@27601
   501
val end_local_theory = transaction (fn _ =>
wenzelm@21294
   502
  (fn Theory (Context.Proof lthy, _) => Theory (Context.Theory (loc_exit lthy), SOME lthy)
wenzelm@21007
   503
    | _ => raise UNDEF));
wenzelm@21007
   504
wenzelm@21007
   505
local
wenzelm@21007
   506
wenzelm@27601
   507
fun local_theory_presentation loc f g = transaction (fn int =>
wenzelm@21294
   508
  (fn Theory (gthy, _) =>
wenzelm@21294
   509
        let
wenzelm@21294
   510
          val finish = loc_finish loc gthy;
wenzelm@25960
   511
          val lthy' = f (loc_begin loc gthy);
wenzelm@21294
   512
        in Theory (finish lthy', SOME lthy') end
wenzelm@27859
   513
    | _ => raise UNDEF) #> tap g);
wenzelm@20963
   514
wenzelm@21007
   515
in
wenzelm@21007
   516
wenzelm@20963
   517
fun local_theory loc f = local_theory_presentation loc f (K I);
wenzelm@20963
   518
fun present_local_theory loc g = local_theory_presentation loc I g;
wenzelm@18955
   519
wenzelm@21007
   520
end;
wenzelm@21007
   521
wenzelm@21007
   522
wenzelm@21007
   523
(* proof transitions *)
wenzelm@21007
   524
wenzelm@27601
   525
fun end_proof f = transaction (fn int =>
wenzelm@24795
   526
  (fn Proof (prf, (finish, _)) =>
wenzelm@27564
   527
        let val state = ProofNode.current prf in
wenzelm@21007
   528
          if can (Proof.assert_bottom true) state then
wenzelm@21007
   529
            let
wenzelm@21007
   530
              val ctxt' = f int state;
wenzelm@21007
   531
              val gthy' = finish ctxt';
wenzelm@21007
   532
            in Theory (gthy', SOME ctxt') end
wenzelm@21007
   533
          else raise UNDEF
wenzelm@21007
   534
        end
wenzelm@27564
   535
    | SkipProof (0, (gthy, _)) => Theory (gthy, NONE)
wenzelm@21007
   536
    | _ => raise UNDEF));
wenzelm@21007
   537
wenzelm@21294
   538
local
wenzelm@21294
   539
wenzelm@27601
   540
fun begin_proof init finish = transaction (fn int =>
wenzelm@21294
   541
  (fn Theory (gthy, _) =>
wenzelm@21294
   542
    let
berghofe@24453
   543
      val prf = init int gthy;
wenzelm@28433
   544
      val skip = ! skip_proofs;
wenzelm@21294
   545
      val schematic = Proof.schematic_goal prf;
wenzelm@21294
   546
    in
wenzelm@28433
   547
      if skip andalso schematic then
wenzelm@21294
   548
        warning "Cannot skip proof of schematic goal statement"
wenzelm@21294
   549
      else ();
wenzelm@28433
   550
      if skip andalso not schematic then
wenzelm@27564
   551
        SkipProof (0, (finish gthy (Proof.global_skip_proof int prf), gthy))
wenzelm@27564
   552
      else Proof (ProofNode.init prf, (finish gthy, gthy))
wenzelm@21294
   553
    end
wenzelm@21294
   554
  | _ => raise UNDEF));
wenzelm@21294
   555
wenzelm@21294
   556
in
wenzelm@21294
   557
wenzelm@24780
   558
fun local_theory_to_proof' loc f = begin_proof
wenzelm@25960
   559
  (fn int => fn gthy => f int (loc_begin loc gthy))
wenzelm@24780
   560
  (loc_finish loc);
wenzelm@24780
   561
berghofe@24453
   562
fun local_theory_to_proof loc f = local_theory_to_proof' loc (K f);
wenzelm@21294
   563
wenzelm@21294
   564
fun theory_to_proof f = begin_proof
wenzelm@24780
   565
  (K (fn Context.Theory thy => f thy | _ => raise UNDEF))
wenzelm@24780
   566
  (K (Context.Theory o ProofContext.theory_of));
wenzelm@21294
   567
wenzelm@21294
   568
end;
wenzelm@21294
   569
wenzelm@27601
   570
val forget_proof = transaction (fn _ =>
wenzelm@21007
   571
  (fn Proof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
wenzelm@21007
   572
    | SkipProof (_, (_, orig_gthy)) => Theory (orig_gthy, NONE)
wenzelm@21007
   573
    | _ => raise UNDEF));
wenzelm@21007
   574
wenzelm@27859
   575
fun present_proof f = transaction (fn _ =>
wenzelm@27564
   576
  (fn Proof (prf, x) => Proof (ProofNode.apply I prf, x)
wenzelm@27564
   577
    | skip as SkipProof _ => skip
wenzelm@27859
   578
    | _ => raise UNDEF) #> tap f);
wenzelm@21177
   579
wenzelm@27601
   580
fun proofs' f = transaction (fn int =>
wenzelm@27564
   581
  (fn Proof (prf, x) => Proof (ProofNode.applys (f int) prf, x)
wenzelm@27564
   582
    | skip as SkipProof _ => skip
wenzelm@16815
   583
    | _ => raise UNDEF));
wenzelm@15668
   584
wenzelm@17904
   585
fun proof' f = proofs' (Seq.single oo f);
wenzelm@17904
   586
val proofs = proofs' o K;
wenzelm@6689
   587
val proof = proof' o K;
wenzelm@16815
   588
wenzelm@27601
   589
fun actual_proof f = transaction (fn _ =>
wenzelm@21007
   590
  (fn Proof (prf, x) => Proof (f prf, x)
wenzelm@20963
   591
    | _ => raise UNDEF));
wenzelm@16815
   592
wenzelm@27601
   593
fun skip_proof f = transaction (fn _ =>
wenzelm@21007
   594
  (fn SkipProof (h, x) => SkipProof (f h, x)
wenzelm@18563
   595
    | _ => raise UNDEF));
wenzelm@18563
   596
wenzelm@27601
   597
fun skip_proof_to_theory pred = transaction (fn _ =>
wenzelm@27564
   598
  (fn SkipProof (d, (gthy, _)) => if pred d then Theory (gthy, NONE) else raise UNDEF
wenzelm@17904
   599
  | _ => raise UNDEF));
wenzelm@5828
   600
wenzelm@5828
   601
wenzelm@5828
   602
wenzelm@5828
   603
(** toplevel transactions **)
wenzelm@5828
   604
wenzelm@27427
   605
(* identification *)
wenzelm@27427
   606
wenzelm@27427
   607
fun get_id (Transition {pos, ...}) = Position.get_id pos;
wenzelm@27427
   608
fun put_id id (tr as Transition {pos, ...}) = position (Position.put_id id pos) tr;
wenzelm@27427
   609
wenzelm@27427
   610
wenzelm@25960
   611
(* thread position *)
wenzelm@25799
   612
wenzelm@25960
   613
fun setmp_thread_position (Transition {pos, ...}) f x =
wenzelm@25819
   614
  Position.setmp_thread_data pos f x;
wenzelm@25799
   615
wenzelm@27606
   616
fun status tr m =
wenzelm@27606
   617
  setmp_thread_position tr (fn () => Output.status (Markup.markup m "")) ();
wenzelm@27606
   618
wenzelm@26602
   619
fun error_msg tr exn_info =
wenzelm@26602
   620
  setmp_thread_position tr (fn () => Output.error_msg (exn_message (EXCURSION_FAIL exn_info))) ();
wenzelm@26602
   621
wenzelm@25799
   622
wenzelm@28095
   623
(* post-transition hooks *)
wenzelm@28095
   624
wenzelm@28103
   625
local val hooks = ref ([]: (transition -> state -> state -> unit) list) in
wenzelm@28095
   626
wenzelm@28095
   627
fun add_hook f = CRITICAL (fn () => change hooks (cons f));
wenzelm@28095
   628
fun get_hooks () = CRITICAL (fn () => ! hooks);
wenzelm@28095
   629
wenzelm@28095
   630
end;
wenzelm@28095
   631
wenzelm@28095
   632
wenzelm@5828
   633
(* apply transitions *)
wenzelm@5828
   634
wenzelm@6664
   635
local
wenzelm@6664
   636
wenzelm@25799
   637
fun app int (tr as Transition {trans, pos, int_only, print, no_timing, ...}) =
wenzelm@25819
   638
  setmp_thread_position tr (fn state =>
wenzelm@25799
   639
    let
wenzelm@26621
   640
      val _ = if not int andalso int_only then warning (command_msg "Interactive-only " tr) else ();
wenzelm@16682
   641
wenzelm@25799
   642
      fun do_timing f x = (warning (command_msg "" tr); timeap f x);
wenzelm@25799
   643
      fun do_profiling f x = profile (! profiling) f x;
wenzelm@16682
   644
wenzelm@26256
   645
      val (result, status) =
wenzelm@25799
   646
         state |> (apply_trans int pos trans
wenzelm@25799
   647
          |> (if ! profiling > 0 andalso not no_timing then do_profiling else I)
wenzelm@25799
   648
          |> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I));
wenzelm@26256
   649
wenzelm@26621
   650
      val _ = if int andalso not (! quiet) andalso print then print_state false result else ();
wenzelm@26256
   651
    in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) status) end);
wenzelm@6664
   652
wenzelm@6664
   653
in
wenzelm@5828
   654
wenzelm@26602
   655
fun transition int tr st =
wenzelm@28095
   656
  let
wenzelm@28095
   657
    val hooks = get_hooks ();
wenzelm@28103
   658
    fun apply_hooks st' = hooks |> List.app (fn f => (try (fn () => f tr st st') (); ()));
wenzelm@28095
   659
wenzelm@28095
   660
    val ctxt = try context_of st;
wenzelm@28095
   661
    val res =
wenzelm@28095
   662
      (case app int tr st of
wenzelm@28095
   663
        (_, SOME TERMINATE) => NONE
wenzelm@28103
   664
      | (st', SOME (EXCURSION_FAIL exn_info)) => SOME (st', SOME exn_info)
wenzelm@28103
   665
      | (st', SOME exn) => SOME (st', SOME (exn_context ctxt exn, at_command tr))
wenzelm@28103
   666
      | (st', NONE) => SOME (st', NONE));
wenzelm@28103
   667
    val _ = (case res of SOME (st', NONE) => apply_hooks st' | _ => ());
wenzelm@28095
   668
  in res end;
wenzelm@6664
   669
wenzelm@6664
   670
end;
wenzelm@5828
   671
wenzelm@5828
   672
wenzelm@27601
   673
(* commit final exit *)
wenzelm@27601
   674
wenzelm@28425
   675
fun commit_exit pos =
wenzelm@27601
   676
  name "end" empty
wenzelm@28425
   677
  |> position pos
wenzelm@27601
   678
  |> add_trans CommitExit
wenzelm@27601
   679
  |> imperative (fn () => warning "Expected to find finished theory");
wenzelm@27601
   680
wenzelm@27601
   681
wenzelm@28425
   682
(* nested commands *)
wenzelm@5828
   683
wenzelm@28425
   684
fun command tr st =
wenzelm@27576
   685
  (case transition (! interact) tr st of
wenzelm@28425
   686
    SOME (st', NONE) => st'
wenzelm@28425
   687
  | SOME (_, SOME exn_info) => raise EXCURSION_FAIL exn_info
wenzelm@27576
   688
  | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr));
wenzelm@27576
   689
wenzelm@28433
   690
wenzelm@28433
   691
(* excursion of units, consisting of commands with proof *)
wenzelm@28433
   692
wenzelm@28425
   693
fun command_result tr st =
wenzelm@28425
   694
  let val st' = command tr st
wenzelm@28425
   695
  in (st', st') end;
wenzelm@17076
   696
wenzelm@28433
   697
fun unit_result do_promise (tr, proof_trs) st =
wenzelm@28433
   698
  let val st' = command tr st in
wenzelm@28433
   699
    if do_promise andalso not (null proof_trs) andalso
wenzelm@28433
   700
      can proof_of st' andalso Proof.can_promise (proof_of st')
wenzelm@28433
   701
    then
wenzelm@28433
   702
      let
wenzelm@28433
   703
        val (body_trs, end_tr) = split_last proof_trs;
wenzelm@28433
   704
        val body_states = ref ([]: state list);
wenzelm@28433
   705
        fun make_result prf =
wenzelm@28433
   706
          let val (states, result_state) =
wenzelm@28433
   707
            (case st' of State (SOME (Proof (_, (_, orig_gthy)), exit), prev)
wenzelm@28433
   708
              => State (SOME (Proof (ProofNode.init prf, (Context.Proof, orig_gthy)), exit), prev))
wenzelm@28433
   709
            |> fold_map command_result body_trs ||> command end_tr
wenzelm@28433
   710
          in body_states := states; context_of result_state end;
wenzelm@28433
   711
        val proof_promise =
wenzelm@28433
   712
          end_tr |> reset_trans |> end_proof (K (Proof.promise_proof make_result));
wenzelm@28433
   713
        val st'' = command proof_promise st';
wenzelm@28433
   714
      in (fn () => (tr, st') :: (body_trs ~~ ! body_states) @ [(end_tr, st'')], st'') end
wenzelm@28433
   715
    else
wenzelm@28433
   716
      let val (states, st'') = fold_map command_result proof_trs st'
wenzelm@28433
   717
      in (fn () => (tr, st') :: (proof_trs ~~ states), st'') end
wenzelm@28433
   718
  end;
wenzelm@28433
   719
wenzelm@28433
   720
fun excursion input =
wenzelm@28425
   721
  let
wenzelm@28433
   722
    val end_pos = if null input then error "No input" else pos_of (fst (List.last input));
wenzelm@28433
   723
wenzelm@28433
   724
    val do_promise = ! future_scheduler andalso Multithreading.max_threads_value () > 1;
wenzelm@28433
   725
    val (future_results, end_state) = fold_map (unit_result do_promise) input toplevel;
wenzelm@28433
   726
    val _ =
wenzelm@28433
   727
      (case end_state of
wenzelm@28433
   728
        State (NONE, SOME (Theory (Context.Theory thy, _), _)) => Thm.join_futures thy
wenzelm@28433
   729
      | _ => error "Unfinished development at end of input");
wenzelm@28433
   730
    val results = maps (fn res => res ()) future_results;
wenzelm@28433
   731
  in (results, fn () => ignore (command (commit_exit end_pos) end_state)) end;
wenzelm@7062
   732
wenzelm@6664
   733
end;