src/Pure/System/isar.ML
author haftmann
Fri, 25 Sep 2009 09:50:31 +0200
changeset 32705 04ce6bb14d85
parent 32486 67972a7f85b7
child 32738 15bb09ca0378
permissions -rw-r--r--
merged
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
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
     4
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
     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
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
27533
85bbd045ac3e added print;
wenzelm
parents: 27530
diff changeset
    14
  val print: unit -> unit
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    15
  val >> : Toplevel.transition -> bool
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    16
  val >>> : Toplevel.transition list -> unit
27529
6a5ccbb1bca0 added Isar.linear_undo;
wenzelm
parents: 27528
diff changeset
    17
  val linear_undo: int -> unit
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
    18
  val undo: int -> unit
27530
df14c9cbd21d export init_point;
wenzelm
parents: 27529
diff changeset
    19
  val kill: unit -> unit
df14c9cbd21d export init_point;
wenzelm
parents: 27529
diff changeset
    20
  val kill_proof: unit -> unit
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    21
  val crashes: exn list ref
26643
99f5407c05ef Isar.toplevel_loop: separate init/welcome flag;
wenzelm
parents: 26606
diff changeset
    22
  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
    23
  val loop: unit -> unit
26606
379596d12f25 replaced Isar loop variants by generic toplevel_loop;
wenzelm
parents: 26605
diff changeset
    24
  val main: unit -> unit
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    25
end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    26
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    27
structure Isar: ISAR =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    28
struct
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    29
27432
c5ec309c6de8 replaced datatype kind by OuterKeyword.category;
wenzelm
parents: 27428
diff changeset
    30
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    31
(** TTY model -- SINGLE-THREADED! **)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    32
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    33
(* the global state *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    34
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    35
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
    36
  (*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
    37
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    38
local
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    39
  val global_history = ref ([]: history);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    40
  val global_state = ref Toplevel.toplevel;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    41
  val global_exn = ref (NONE: (exn * string) option);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    42
in
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    43
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    44
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
    45
  let
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    46
    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
    47
      | 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
    48
  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
    49
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    50
fun state () = NAMED_CRITICAL "Isar" (fn () => ! global_state);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    51
fun set_state state = NAMED_CRITICAL "Isar" (fn () => global_state := state);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    52
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    53
fun exn () = NAMED_CRITICAL "Isar" (fn () => ! global_exn);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    54
fun set_exn exn =  NAMED_CRITICAL "Isar" (fn () => global_exn := exn);
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
end;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    57
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    58
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    59
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
    60
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    61
fun context () = Toplevel.context_of (state ())
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    62
  handle Toplevel.UNDEF => error "Unknown context";
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
fun goal () = #2 (#2 (Proof.get_goal (Toplevel.proof_of (state ()))))
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    65
  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
    66
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    67
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
    68
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    69
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    70
(* history navigation *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    71
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    72
local
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
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
    75
  | find_and_undo which ((prev, tr) :: hist) =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    76
      ((case Toplevel.init_of tr of SOME name => ThyInfo.kill_thy name | NONE => ());
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    77
        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
    78
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    79
in
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 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
    82
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    83
fun undo n = edit_history n (fn st => fn hist =>
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    84
  find_and_undo (if Toplevel.is_proof st then K true else OuterKeyword.is_theory) hist);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    85
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    86
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
    87
  find_and_undo
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    88
    (if Toplevel.is_proof st then OuterKeyword.is_theory else OuterKeyword.is_theory_begin) hist);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    89
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    90
fun kill_proof () = 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
    91
  if Toplevel.is_proof st then find_and_undo OuterKeyword.is_theory hist
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    92
  else raise Toplevel.UNDEF);
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
end;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    95
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    96
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    97
(* interactive state transformations *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    98
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    99
fun op >> tr =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   100
  (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
   101
    NONE => false
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   102
  | 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
   103
  | SOME (st', NONE) =>
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   104
      let
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   105
        val name = Toplevel.name_of tr;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   106
        val _ = if OuterKeyword.is_theory_begin name then init () else ();
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   107
        val _ =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   108
          if OuterKeyword.is_regular name
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   109
          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
   110
      in true end);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   111
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   112
fun op >>> [] = ()
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   113
  | 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
   114
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   115
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   116
(* toplevel loop *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   117
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   118
val crashes = ref ([]: exn list);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   119
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   120
local
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   121
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   122
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
   123
  let
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   124
    fun check_secure () =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   125
      (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
   126
  in
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   127
    (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
   128
      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
   129
    | SOME (tr, src') => if op >> tr orelse check_secure () then raw_loop secure src' else ())
29370
wenzelm
parents: 29348
diff changeset
   130
    handle exn =>
31478
5e412e4c6546 ML_Compiler.exn_message;
wenzelm
parents: 31443
diff changeset
   131
      (Output.error_msg (ML_Compiler.exn_message exn)
29370
wenzelm
parents: 29348
diff changeset
   132
        handle crash =>
wenzelm
parents: 29348
diff changeset
   133
          (CRITICAL (fn () => change crashes (cons crash));
wenzelm
parents: 29348
diff changeset
   134
            warning "Recovering from Isar toplevel crash -- see also Isar.crashes");
wenzelm
parents: 29348
diff changeset
   135
          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
   136
  end;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   137
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   138
in
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   139
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   140
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
   141
 (Context.set_thread_data NONE;
32486
67972a7f85b7 removed old Isar document model;
wenzelm
parents: 31478
diff changeset
   142
  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
   143
  if welcome then writeln (Session.welcome ()) else ();
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   144
  uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ());
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   145
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   146
end;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   147
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   148
fun loop () =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   149
  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
   150
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   151
fun main () =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   152
  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
   153
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   154
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   155
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
   156
(** command syntax **)
28300
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
local
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
structure P = OuterParse and K = OuterKeyword;
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   161
val op >> = Scan.>>;
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   162
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   163
in
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   164
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
   165
(* 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
   166
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
   167
val _ =
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
   168
  OuterSyntax.improper_command "init_toplevel" "init toplevel point-of-interest" K.control
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
    (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
   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 _ =
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
   172
  OuterSyntax.improper_command "linear_undo" "undo commands" K.control
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
   173
    (Scan.optional P.nat 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
   174
      (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
   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 _ =
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
   177
  OuterSyntax.improper_command "undo" "undo commands (skipping closed proofs)" K.control
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
   178
    (Scan.optional P.nat 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
   179
      (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
   180
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
val _ =
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
  OuterSyntax.improper_command "undos_proof" "undo commands (skipping closed proofs)" K.control
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
    (Scan.optional P.nat 1 >> (fn n => Toplevel.no_timing o
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
   184
      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
   185
        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
   186
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
val _ =
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
  OuterSyntax.improper_command "cannot_undo" "partial undo -- Proof General legacy" K.control
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
    (P.name >>
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
      (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
   191
        | 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
   192
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
val _ =
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
   194
  OuterSyntax.improper_command "kill" "kill partial proof or theory development" K.control
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
   195
    (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
   196
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   197
end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   198
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   199
end;