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