src/Pure/Isar/isar.ML
author wenzelm
Thu, 10 Jul 2008 17:43:02 +0200
changeset 27527 02b2c9eab4e2
parent 27524 cd95da386e9a
child 27528 4bbf70c35a6e
permissions -rw-r--r--
tty interaction: do not move point after error;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/isar.ML
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     4
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     5
The global Isabelle/Isar state and main read-eval-print loop.
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     6
*)
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     7
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     8
signature ISAR =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     9
sig
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    10
  val state: unit -> Toplevel.state
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    11
  val exn: unit -> (exn * string) option
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    12
  val context: unit -> Proof.context
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    13
  val goal: unit -> thm
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    14
  val >> : Toplevel.transition -> bool
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    15
  val >>> : Toplevel.transition list -> unit
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    16
  val undo: int -> unit
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    17
  val crashes: exn list ref
26643
99f5407c05ef Isar.toplevel_loop: separate init/welcome flag;
wenzelm
parents: 26606
diff changeset
    18
  val toplevel_loop: {init: bool, welcome: bool, sync: bool, secure: bool} -> unit
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    19
  val loop: unit -> unit
26606
379596d12f25 replaced Isar loop variants by generic toplevel_loop;
wenzelm
parents: 26605
diff changeset
    20
  val main: unit -> unit
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    21
end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    22
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    23
structure Isar: ISAR =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    24
struct
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    25
27432
c5ec309c6de8 replaced datatype kind by OuterKeyword.category;
wenzelm
parents: 27428
diff changeset
    26
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    27
(** individual toplevel commands **)
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    28
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    29
(* unique identification *)
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    30
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    31
type id = string;
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    32
val no_id : id = "";
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    33
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    34
fun identify tr =
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    35
  (case Toplevel.get_id tr of
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    36
    SOME id => (id, tr)
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    37
  | NONE =>
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    38
      let val id = "isabelle:" ^ serial_string ()
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    39
      in (id, Toplevel.put_id id tr) end);
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    40
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    41
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    42
(* command category *)
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    43
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    44
datatype category = Empty | Theory | Proof | Diag | Control;
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    45
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    46
fun category_of tr =
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    47
  let val name = Toplevel.name_of tr in
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    48
    if name = "" then Empty
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    49
    else if OuterKeyword.is_theory name then Theory
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    50
    else if OuterKeyword.is_proof name then Proof
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    51
    else if OuterKeyword.is_diag name then Diag
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    52
    else Control
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    53
  end;
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    54
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    55
val is_theory = fn Theory => true | _ => false;
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    56
val is_proper = fn Theory => true | Proof => true | _ => false;
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    57
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    58
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    59
(* datatype command *)
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    60
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    61
datatype status =
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    62
  Initial |
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
    63
  Result of Toplevel.state * (exn * string) option;
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    64
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    65
datatype command = Command of
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    66
 {category: category,
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    67
  transition: Toplevel.transition,
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    68
  status: status};
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    69
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    70
fun make_command (category, transition, status) =
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    71
  Command {category = category, transition = transition, status = status};
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    72
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
    73
val empty_command =
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
    74
  make_command (Empty, Toplevel.empty, Result (Toplevel.toplevel, NONE));
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    75
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    76
fun map_command f (Command {category, transition, status}) =
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    77
  make_command (f (category, transition, status));
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    78
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    79
fun map_status f = map_command (fn (category, transition, status) =>
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    80
  (category, transition, f status));
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    81
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    82
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    83
(* global collection of identified commands *)
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    84
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
    85
fun err_dup id = sys_error ("Duplicate command " ^ quote id);
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
    86
fun err_undef id = sys_error ("Unknown command " ^ quote id);
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
    87
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    88
local
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    89
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    90
val empty_commands = Graph.empty: command Graph.T;
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    91
val global_commands = ref empty_commands;
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    92
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    93
in
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    94
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    95
fun change_commands f = NAMED_CRITICAL "Isar" (fn () => change global_commands f)
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
    96
  handle Graph.DUP id => err_dup id | Graph.UNDEF id => err_undef id;
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
    97
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    98
fun init_commands () = change_commands (K empty_commands);
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
    99
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   100
fun the_command id =
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   101
  let val Command cmd =
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   102
    if id = no_id then empty_command
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   103
    else (Graph.get_node (! global_commands) id handle Graph.UNDEF _ => err_undef id)
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   104
  in cmd end;
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   105
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   106
fun prev_command id =
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   107
  if id = no_id then NONE
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   108
  else
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   109
    (case Graph.imm_preds (! global_commands) id handle Graph.UNDEF _ => err_undef id of
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   110
      [] => NONE
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   111
    | [prev] => SOME prev
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   112
    | _ => sys_error ("Non-linear command dependency " ^ quote id));
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   113
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   114
end;
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   115
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   116
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   117
fun the_result id =
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   118
  (case the_command id of
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   119
    {status = Result res, ...} => res
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   120
  | {transition, ...} => error ("Unfinished command " ^ Toplevel.str_of transition));
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   121
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   122
val the_state = #1 o the_result;
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   123
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   124
fun new_command prev (id, cmd) =
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   125
  change_commands (Graph.new_node (id, cmd) #> prev <> no_id ? Graph.add_edge (prev, id));
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   126
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   127
fun dispose_command id = change_commands (Graph.del_nodes [id]);
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   128
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   129
fun change_command_status id f = change_commands (Graph.map_node id (map_status f));
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   130
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   131
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   132
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   133
(** TTY interaction **)
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   134
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   135
(* global point *)
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   136
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   137
local val global_point = ref no_id in
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   138
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   139
fun change_point f = NAMED_CRITICAL "Isar" (fn () => change global_point f);
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   140
fun point () = NAMED_CRITICAL "Isar" (fn () => ! global_point);
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   141
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   142
end;
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   143
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   144
fun set_point id = change_point (K id);
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   145
fun init_point () = set_point no_id;
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   146
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   147
fun point_result () = NAMED_CRITICAL "Isar" (fn () =>
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   148
  let val id = point () in (id, the_result id) end);
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   149
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   150
fun state () = #1 (#2 (point_result ()));
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   151
fun exn () = #2 (#2 (point_result ()));
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   152
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   153
fun context () =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   154
  Toplevel.context_of (state ())
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   155
    handle Toplevel.UNDEF => error "Unknown context";
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   156
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   157
fun goal () =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   158
  #2 (#2 (Proof.get_goal (Toplevel.proof_of (state ()))))
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   159
    handle Toplevel.UNDEF => error "No goal present";
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   160
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   161
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   162
(* interactive state transformations --- NOT THREAD-SAFE! *)
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   163
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   164
nonfix >> >>>;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   165
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   166
fun >> raw_tr =
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   167
  let
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   168
    val (id, tr) = identify raw_tr;
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   169
    val (prev, (prev_state, _)) = point_result ();
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   170
    val category = category_of tr;
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   171
    val _ = new_command prev (id, make_command (category, tr, Initial));
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   172
  in
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   173
    (case Toplevel.transition true tr prev_state of
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   174
      NONE => (dispose_command id; false)
27527
02b2c9eab4e2 tty interaction: do not move point after error;
wenzelm
parents: 27524
diff changeset
   175
    | SOME (result as (_, err)) =>
02b2c9eab4e2 tty interaction: do not move point after error;
wenzelm
parents: 27524
diff changeset
   176
        (change_command_status id (K (Result result));
02b2c9eab4e2 tty interaction: do not move point after error;
wenzelm
parents: 27524
diff changeset
   177
          Option.map (Toplevel.error_msg tr) err;
02b2c9eab4e2 tty interaction: do not move point after error;
wenzelm
parents: 27524
diff changeset
   178
          if is_some err orelse category = Control then dispose_command id
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   179
          else set_point id;
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   180
          true))
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   181
  end;
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   182
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   183
fun >>> [] = ()
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   184
  | >>> (tr :: trs) = if >> tr then >>> trs else ();
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   185
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   186
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   187
(* old-style navigation *)
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   188
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   189
local
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   190
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   191
fun err_undo () = error "Undo history exhausted";
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   192
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   193
fun get_prev id = the_default no_id (prev_command id);
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   194
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   195
fun find_command pred id =
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   196
  (case #category (the_command id) of
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   197
    Empty => err_undo ()
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   198
  | category => if pred category then id else find_command pred (get_prev id));
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   199
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   200
fun undo_proper id =
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   201
  let
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   202
    val base = find_command (if Toplevel.is_proof (the_state id) then is_proper else is_theory) id;
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   203
    val {category, transition, ...} = the_command base;
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   204
  in
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   205
    (case Toplevel.init_of transition of
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   206
      SOME name => get_prev base before ThyInfo.kill_thy name
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   207
    | NONE => get_prev base)
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   208
  end;
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   209
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   210
in
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   211
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   212
fun undo n = change_point (funpow n undo_proper);
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   213
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   214
end;
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   215
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   216
26606
379596d12f25 replaced Isar loop variants by generic toplevel_loop;
wenzelm
parents: 26605
diff changeset
   217
(* toplevel loop *)
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   218
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   219
val crashes = ref ([]: exn list);
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   220
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   221
local
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   222
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   223
fun raw_loop secure src =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   224
  let
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   225
    fun check_secure () =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   226
      (if secure then warning "Secure loop -- cannot exit to ML" else (); secure);
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   227
    val prev = point ();
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   228
    val prev_name = Toplevel.name_of (#transition (the_command prev));
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   229
    val prompt_markup =
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   230
      prev <> no_id ? Markup.markup
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   231
        (Markup.properties [(Markup.idN, prev), (Markup.nameN, prev_name)] Markup.prompt);
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   232
  in
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   233
    (case Source.get_single (Source.set_prompt (prompt_markup Source.default_prompt) src) of
26606
379596d12f25 replaced Isar loop variants by generic toplevel_loop;
wenzelm
parents: 26605
diff changeset
   234
      NONE => if secure then quit () else ()
379596d12f25 replaced Isar loop variants by generic toplevel_loop;
wenzelm
parents: 26605
diff changeset
   235
    | SOME (tr, src') => if >> tr orelse check_secure () then raw_loop secure src' else ())
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   236
    handle exn => (Output.error_msg (Toplevel.exn_message exn) handle crash =>
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   237
      (CRITICAL (fn () => change crashes (cons crash));
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   238
        warning "Recovering after Isar toplevel crash -- see also Isar.crashes");
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   239
      raw_loop secure src)
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   240
  end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   241
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   242
in
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   243
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   244
fun toplevel_loop {init, welcome, sync, secure} =
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   245
 (Context.set_thread_data NONE;
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   246
  if init then (init_point (); init_commands ()) else ();
26643
99f5407c05ef Isar.toplevel_loop: separate init/welcome flag;
wenzelm
parents: 26606
diff changeset
   247
  if welcome then writeln (Session.welcome ()) else ();
26606
379596d12f25 replaced Isar loop variants by generic toplevel_loop;
wenzelm
parents: 26605
diff changeset
   248
  uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ());
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   249
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   250
end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   251
26643
99f5407c05ef Isar.toplevel_loop: separate init/welcome flag;
wenzelm
parents: 26606
diff changeset
   252
fun loop () =
99f5407c05ef Isar.toplevel_loop: separate init/welcome flag;
wenzelm
parents: 26606
diff changeset
   253
  toplevel_loop {init = false, welcome = false, sync = false, secure = Secure.is_secure ()};
99f5407c05ef Isar.toplevel_loop: separate init/welcome flag;
wenzelm
parents: 26606
diff changeset
   254
fun main () =
99f5407c05ef Isar.toplevel_loop: separate init/welcome flag;
wenzelm
parents: 26606
diff changeset
   255
  toplevel_loop {init = true, welcome = true, sync = false, secure = Secure.is_secure ()};
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   256
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   257
end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   258