src/Pure/System/isar.ML
author wenzelm
Thu, 03 Jun 2010 22:31:59 +0200
changeset 37306 2bde06a2a706
parent 37239 54b444874be1
child 38138 dc65ed8bbb43
permissions -rw-r--r--
discontinued obsolete Isar.context() -- long superseded by @{context};
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30175
62ba490670e8 fixed headers;
wenzelm
parents: 30173
diff changeset
     1
(*  Title:      Pure/System/isar.ML
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     3
37306
2bde06a2a706 discontinued obsolete Isar.context() -- long superseded by @{context};
wenzelm
parents: 37239
diff changeset
     4
Global state of the raw Isar read-eval-print loop.
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     5
*)
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
signature ISAR =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     8
sig
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
     9
  val init: unit -> unit
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    10
  val exn: unit -> (exn * string) option
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    11
  val state: unit -> Toplevel.state
33289
d0c2ef490613 Isar.goal: Proof.simple_goal, not raw version;
wenzelm
parents: 33223
diff changeset
    12
  val goal: unit -> {context: Proof.context, facts: thm list, goal: thm}
27533
85bbd045ac3e added print;
wenzelm
parents: 27530
diff changeset
    13
  val print: unit -> unit
26605
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
27529
6a5ccbb1bca0 added Isar.linear_undo;
wenzelm
parents: 27528
diff changeset
    16
  val linear_undo: int -> unit
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    17
  val undo: int -> unit
27530
df14c9cbd21d export init_point;
wenzelm
parents: 27529
diff changeset
    18
  val kill: unit -> unit
df14c9cbd21d export init_point;
wenzelm
parents: 27529
diff changeset
    19
  val kill_proof: unit -> unit
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32486
diff changeset
    20
  val crashes: exn list Unsynchronized.ref
26643
99f5407c05ef Isar.toplevel_loop: separate init/welcome flag;
wenzelm
parents: 26606
diff changeset
    21
  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
    22
  val loop: unit -> unit
26606
379596d12f25 replaced Isar loop variants by generic toplevel_loop;
wenzelm
parents: 26605
diff changeset
    23
  val main: unit -> unit
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    24
end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    25
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    26
structure Isar: ISAR =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    27
struct
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    28
27432
c5ec309c6de8 replaced datatype kind by OuterKeyword.category;
wenzelm
parents: 27428
diff changeset
    29
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    30
(** TTY model -- SINGLE-THREADED! **)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    31
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    32
(* the global state *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    33
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    34
type history = (Toplevel.state * Toplevel.transition) list;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    35
  (*previous state, state transition -- regular commands only*)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    36
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    37
local
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32486
diff changeset
    38
  val global_history = Unsynchronized.ref ([]: history);
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32486
diff changeset
    39
  val global_state = Unsynchronized.ref Toplevel.toplevel;
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32486
diff changeset
    40
  val global_exn = Unsynchronized.ref (NONE: (exn * string) option);
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    41
in
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    42
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    43
fun edit_history count f = NAMED_CRITICAL "Isar" (fn () =>
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    44
  let
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    45
    fun edit 0 (st, hist) = (global_history := hist; global_state := st; global_exn := NONE)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    46
      | edit n (st, hist) = edit (n - 1) (f st hist);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    47
  in edit count (! global_state, ! global_history) end);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    48
33223
d27956b4d3b4 non-critical atomic accesses;
wenzelm
parents: 32786
diff changeset
    49
fun state () = ! global_state;
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    50
33223
d27956b4d3b4 non-critical atomic accesses;
wenzelm
parents: 32786
diff changeset
    51
fun exn () = ! global_exn;
d27956b4d3b4 non-critical atomic accesses;
wenzelm
parents: 32786
diff changeset
    52
fun set_exn exn =  global_exn := exn;
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    53
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    54
end;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    55
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    56
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    57
fun init () = edit_history 1 (K (K (Toplevel.toplevel, [])));
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    58
33289
d0c2ef490613 Isar.goal: Proof.simple_goal, not raw version;
wenzelm
parents: 33223
diff changeset
    59
fun goal () = Proof.goal (Toplevel.proof_of (state ()))
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    60
  handle Toplevel.UNDEF => error "No goal present";
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    61
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    62
fun print () = Toplevel.print_state false (state ());
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    63
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    64
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    65
(* history navigation *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    66
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    67
local
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    68
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    69
fun find_and_undo _ [] = error "Undo history exhausted"
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    70
  | find_and_undo which ((prev, tr) :: hist) =
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36953
diff changeset
    71
      ((case Toplevel.init_of tr of SOME name => Thy_Info.kill_thy name | NONE => ());
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    72
        if which (Toplevel.name_of tr) then (prev, hist) else find_and_undo which hist);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    73
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    74
in
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    75
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    76
fun linear_undo n = edit_history n (K (find_and_undo (K true)));
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    77
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    78
fun undo n = edit_history n (fn st => fn hist =>
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
    79
  find_and_undo (if Toplevel.is_proof st then K true else Keyword.is_theory) hist);
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    80
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    81
fun kill () = edit_history 1 (fn st => fn hist =>
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    82
  find_and_undo
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
    83
    (if Toplevel.is_proof st then Keyword.is_theory else Keyword.is_theory_begin) hist);
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    84
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    85
fun kill_proof () = edit_history 1 (fn st => fn hist =>
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
    86
  if Toplevel.is_proof st then find_and_undo Keyword.is_theory hist
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    87
  else raise Toplevel.UNDEF);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    88
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    89
end;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    90
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    91
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    92
(* interactive state transformations *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    93
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    94
fun op >> tr =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    95
  (case Toplevel.transition true tr (state ()) of
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    96
    NONE => false
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    97
  | SOME (_, SOME err) => (set_exn (SOME err); Toplevel.error_msg tr err; true)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    98
  | SOME (st', NONE) =>
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    99
      let
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   100
        val name = Toplevel.name_of tr;
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
   101
        val _ = if Keyword.is_theory_begin name then init () else ();
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   102
        val _ =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
   103
          if Keyword.is_regular name
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   104
          then edit_history 1 (fn st => fn hist => (st', (st, tr) :: hist)) else ();
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   105
      in true end);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   106
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   107
fun op >>> [] = ()
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   108
  | op >>> (tr :: trs) = if op >> tr then op >>> trs else ();
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   109
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   110
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   111
(* toplevel loop *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   112
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32486
diff changeset
   113
val crashes = Unsynchronized.ref ([]: exn list);
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   114
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   115
local
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   116
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   117
fun raw_loop secure src =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   118
  let
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   119
    fun check_secure () =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   120
      (if secure then warning "Secure loop -- cannot exit to ML" else (); secure);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   121
  in
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   122
    (case Source.get_single (Source.set_prompt Source.default_prompt src) of
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   123
      NONE => if secure then quit () else ()
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   124
    | SOME (tr, src') => if op >> tr orelse check_secure () then raw_loop secure src' else ())
29370
wenzelm
parents: 29348
diff changeset
   125
    handle exn =>
31478
5e412e4c6546 ML_Compiler.exn_message;
wenzelm
parents: 31443
diff changeset
   126
      (Output.error_msg (ML_Compiler.exn_message exn)
29370
wenzelm
parents: 29348
diff changeset
   127
        handle crash =>
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32486
diff changeset
   128
          (CRITICAL (fn () => Unsynchronized.change crashes (cons crash));
29370
wenzelm
parents: 29348
diff changeset
   129
            warning "Recovering from Isar toplevel crash -- see also Isar.crashes");
wenzelm
parents: 29348
diff changeset
   130
          raw_loop secure src)
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   131
  end;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   132
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   133
in
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   134
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   135
fun toplevel_loop {init = do_init, welcome, sync, secure} =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   136
 (Context.set_thread_data NONE;
37239
54b444874be1 uniform ML environment setup for Isar and PG;
wenzelm
parents: 37216
diff changeset
   137
  Secure.Isar_setup ();
32486
67972a7f85b7 removed old Isar document model;
wenzelm
parents: 31478
diff changeset
   138
  if do_init then init () else ();
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   139
  if welcome then writeln (Session.welcome ()) else ();
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
   140
  uninterruptible (fn _ => fn () => raw_loop secure (Outer_Syntax.isar sync)) ());
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   141
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   142
end;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   143
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   144
fun loop () =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   145
  toplevel_loop {init = false, welcome = false, sync = false, secure = Secure.is_secure ()};
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   146
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   147
fun main () =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   148
  toplevel_loop {init = true, welcome = true, sync = false, secure = Secure.is_secure ()};
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   149
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   150
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   151
30173
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   152
(** command syntax **)
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   153
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   154
local
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   155
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   156
val op >> = Scan.>>;
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   157
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   158
in
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   159
30173
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   160
(* global history *)
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   161
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   162
val _ =
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
   163
  Outer_Syntax.improper_command "init_toplevel" "init toplevel point-of-interest" Keyword.control
30173
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   164
    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative init));
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   165
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   166
val _ =
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
   167
  Outer_Syntax.improper_command "linear_undo" "undo commands" Keyword.control
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
   168
    (Scan.optional Parse.nat 1 >>
30173
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   169
      (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => linear_undo n)));
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   170
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   171
val _ =
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
   172
  Outer_Syntax.improper_command "undo" "undo commands (skipping closed proofs)" Keyword.control
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
   173
    (Scan.optional Parse.nat 1 >>
30173
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   174
      (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => undo n)));
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   175
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   176
val _ =
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
   177
  Outer_Syntax.improper_command "undos_proof" "undo commands (skipping closed proofs)"
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
   178
    Keyword.control
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
   179
    (Scan.optional Parse.nat 1 >> (fn n => Toplevel.no_timing o
30173
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   180
      Toplevel.keep (fn state =>
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   181
        if Toplevel.is_proof state then (undo n; print ()) else raise Toplevel.UNDEF)));
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   182
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   183
val _ =
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
   184
  Outer_Syntax.improper_command "cannot_undo" "partial undo -- Proof General legacy"
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
   185
    Keyword.control
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33289
diff changeset
   186
    (Parse.name >>
30173
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   187
      (fn "end" => Toplevel.no_timing o Toplevel.imperative (fn () => undo 1)
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   188
        | txt => Toplevel.imperative (fn () => error ("Cannot undo " ^ quote txt))));
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   189
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   190
val _ =
36953
2af1ad9aa1a3 renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
wenzelm
parents: 36950
diff changeset
   191
  Outer_Syntax.improper_command "kill" "kill partial proof or theory development" Keyword.control
30173
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   192
    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill));
eabece26b89b moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
wenzelm
parents: 29417
diff changeset
   193
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   194
end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   195
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   196
end;