src/Pure/Isar/isar.ML
author wenzelm
Fri, 09 Jan 2009 23:33:59 +0100
changeset 29417 779ff1187327
parent 29370 98aaf2cd873f
permissions -rw-r--r--
added running task markup;
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
    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
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    25
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    26
  type id = string
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    27
  val no_id: id
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    28
  val create_command: Toplevel.transition -> id
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
    29
  val insert_command: id -> id -> unit
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
    30
  val remove_command: id -> unit
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    31
end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    32
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    33
structure Isar: ISAR =
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    34
struct
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
    35
27432
c5ec309c6de8 replaced datatype kind by OuterKeyword.category;
wenzelm
parents: 27428
diff changeset
    36
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    37
(** TTY model -- SINGLE-THREADED! **)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    38
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    39
(* the global state *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    40
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    41
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
    42
  (*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
    43
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    44
local
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    45
  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
    46
  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
    47
  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
    48
in
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 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
    51
  let
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    52
    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
    53
      | 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
    54
  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
    55
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    56
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
    57
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
    58
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    59
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
    60
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
    61
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    62
end;
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
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
    66
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    67
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
    68
  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
    69
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    70
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
    71
  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
    72
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    73
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
    74
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
(* history navigation *)
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
local
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    79
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    80
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
    81
  | 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
    82
      ((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
    83
        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
    84
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    85
in
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    86
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    87
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
    88
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    89
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
    90
  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
    91
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    92
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
    93
  find_and_undo
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    94
    (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
    95
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    96
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
    97
  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
    98
  else raise Toplevel.UNDEF);
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
    99
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   100
end;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   101
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   102
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   103
(* interactive state transformations *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   104
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   105
fun op >> tr =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   106
  (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
   107
    NONE => false
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   108
  | 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
   109
  | SOME (st', NONE) =>
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   110
      let
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   111
        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
   112
        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
   113
        val _ =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   114
          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
   115
          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
   116
      in true end);
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
fun op >>> [] = ()
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   119
  | 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
   120
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
(* toplevel loop *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   123
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   124
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
   125
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   126
local
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   127
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   128
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
   129
  let
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   130
    fun check_secure () =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   131
      (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
   132
  in
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   133
    (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
   134
      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
   135
    | SOME (tr, src') => if op >> tr orelse check_secure () then raw_loop secure src' else ())
29370
wenzelm
parents: 29348
diff changeset
   136
    handle exn =>
wenzelm
parents: 29348
diff changeset
   137
      (Output.error_msg (Toplevel.exn_message exn)
wenzelm
parents: 29348
diff changeset
   138
        handle crash =>
wenzelm
parents: 29348
diff changeset
   139
          (CRITICAL (fn () => change crashes (cons crash));
wenzelm
parents: 29348
diff changeset
   140
            warning "Recovering from Isar toplevel crash -- see also Isar.crashes");
wenzelm
parents: 29348
diff changeset
   141
          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
   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
in
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
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
   147
 (Context.set_thread_data NONE;
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   148
  if do_init then init () else ();  (* FIXME init editor model *)
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   149
  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
   150
  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
   151
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   152
end;
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
fun loop () =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   155
  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
   156
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   157
fun main () =
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   158
  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
   159
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   160
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   161
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   162
(** individual toplevel commands **)
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   163
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   164
(* unique identification *)
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   165
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   166
type id = string;
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   167
val no_id : id = "";
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   168
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   169
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   170
(* command category *)
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   171
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   172
datatype category = Empty | Theory | Proof | Diag | Control;
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   173
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   174
fun category_of tr =
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   175
  let val name = Toplevel.name_of tr in
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   176
    if name = "" then Empty
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   177
    else if OuterKeyword.is_theory name then Theory
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   178
    else if OuterKeyword.is_proof name then Proof
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   179
    else if OuterKeyword.is_diag name then Diag
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   180
    else Control
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   181
  end;
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   182
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   183
val is_theory = fn Theory => true | _ => false;
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   184
val is_proper = fn Theory => true | Proof => true | _ => false;
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   185
val is_regular = fn Control => false | _ => true;
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   186
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   187
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   188
(* command status *)
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   189
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   190
datatype status =
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   191
  Unprocessed |
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   192
  Running |
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   193
  Failed of exn * string |
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   194
  Finished of Toplevel.state;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   195
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   196
fun status_markup Unprocessed = Markup.unprocessed
29417
779ff1187327 added running task markup;
wenzelm
parents: 29370
diff changeset
   197
  | status_markup Running = (Markup.runningN, [])
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   198
  | status_markup (Failed _) = Markup.failed
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   199
  | status_markup (Finished _) = Markup.finished;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   200
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   201
fun run int tr state =
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   202
  (case Toplevel.transition int tr state of
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   203
    NONE => NONE
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   204
  | SOME (_, SOME err) => (Toplevel.error_msg tr err; SOME (Failed err))
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   205
  | SOME (state', NONE) => SOME (Finished state'));
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   206
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   207
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   208
(* datatype command *)
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   209
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   210
datatype command = Command of
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   211
 {category: category,
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   212
  transition: Toplevel.transition,
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   213
  status: status};
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   214
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   215
fun make_command (category, transition, status) =
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   216
  Command {category = category, transition = transition, status = status};
27438
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   217
9b2427cc234e command: always keep transition, not just as initial status;
wenzelm
parents: 27432
diff changeset
   218
val empty_command =
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   219
  make_command (Empty, Toplevel.empty, Finished Toplevel.toplevel);
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   220
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   221
fun map_command f (Command {category, transition, status}) =
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   222
  make_command (f (category, transition, status));
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   223
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   224
fun map_status f = map_command (fn (category, transition, status) =>
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   225
  (category, transition, f status));
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   226
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   227
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   228
(* global collection of identified commands *)
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   229
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   230
fun err_dup id = sys_error ("Duplicate command " ^ quote id);
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   231
fun err_undef id = sys_error ("Unknown command " ^ quote id);
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   232
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   233
local val global_commands = ref (Graph.empty: command Graph.T) in
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   234
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   235
fun change_commands f = NAMED_CRITICAL "Isar" (fn () => change global_commands f)
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   236
  handle Graph.DUP bad => err_dup bad | Graph.UNDEF bad => err_undef bad;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   237
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   238
fun get_commands () = NAMED_CRITICAL "Isar" (fn () => ! global_commands);
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   239
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   240
end;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   241
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   242
fun add_edge (id1, id2) =
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   243
  if id1 = no_id orelse id2 = no_id then I else Graph.add_edge (id1, id2);
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   244
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   245
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   246
fun init_commands () = change_commands (K Graph.empty);
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   247
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   248
fun the_command id =
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   249
  let val Command cmd =
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   250
    if id = no_id then empty_command
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   251
    else (Graph.get_node (get_commands ()) id handle Graph.UNDEF bad => err_undef bad)
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   252
  in cmd end;
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   253
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   254
fun prev_command id =
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   255
  if id = no_id then no_id
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   256
  else
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   257
    (case Graph.imm_preds (get_commands ()) id handle Graph.UNDEF bad => err_undef bad of
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   258
      [] => no_id
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   259
    | [prev] => prev
27518
1e5f48e01e5e misc tuning;
wenzelm
parents: 27501
diff changeset
   260
    | _ => sys_error ("Non-linear command dependency " ^ quote id));
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   261
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   262
fun next_commands id =
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   263
  if id = no_id then []
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   264
  else Graph.imm_succs (get_commands ()) id handle Graph.UNDEF bad => err_undef bad;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   265
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   266
fun descendant_commands ids =
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   267
  Graph.all_succs (get_commands ()) (distinct (op =) (filter_out (fn id => id = no_id) ids))
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   268
    handle Graph.UNDEF bad => err_undef bad;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   269
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   270
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   271
(* maintain status *)
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   272
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   273
fun report_status markup id = Toplevel.status (#transition (the_command id)) markup;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   274
27662
34e7236443f3 tty loop: do not report status;
wenzelm
parents: 27622
diff changeset
   275
fun update_status status id = change_commands (Graph.map_node id (map_status (K status)));
34e7236443f3 tty loop: do not report status;
wenzelm
parents: 27622
diff changeset
   276
34e7236443f3 tty loop: do not report status;
wenzelm
parents: 27622
diff changeset
   277
fun report_update_status status id =
34e7236443f3 tty loop: do not report status;
wenzelm
parents: 27622
diff changeset
   278
  change_commands (Graph.map_node id (map_status (fn old_status =>
34e7236443f3 tty loop: do not report status;
wenzelm
parents: 27622
diff changeset
   279
    let val markup = status_markup status
34e7236443f3 tty loop: do not report status;
wenzelm
parents: 27622
diff changeset
   280
    in if markup <> status_markup old_status then report_status markup id else (); status end)));
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   281
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   282
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   283
(* create and dispose commands *)
27524
cd95da386e9a provide old-style undo operation (still unused);
wenzelm
parents: 27518
diff changeset
   284
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   285
fun create_command raw_tr =
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   286
  let
29314
wenzelm
parents: 29313
diff changeset
   287
    val (id, tr) =
wenzelm
parents: 29313
diff changeset
   288
      (case Toplevel.get_id raw_tr of
wenzelm
parents: 29313
diff changeset
   289
        SOME id => (id, raw_tr)
wenzelm
parents: 29313
diff changeset
   290
      | NONE =>
wenzelm
parents: 29313
diff changeset
   291
          let val id =
wenzelm
parents: 29313
diff changeset
   292
            if ! Toplevel.debug then "isabelle:" ^ Toplevel.name_of raw_tr ^ serial_string ()
wenzelm
parents: 29313
diff changeset
   293
            else "isabelle:" ^ serial_string ()
wenzelm
parents: 29313
diff changeset
   294
          in (id, Toplevel.put_id id raw_tr) end);
wenzelm
parents: 29313
diff changeset
   295
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   296
    val cmd = make_command (category_of tr, tr, Unprocessed);
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   297
    val _ = change_commands (Graph.new_node (id, cmd));
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   298
  in id end;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   299
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   300
fun dispose_commands ids =
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   301
  let
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   302
    val desc = descendant_commands ids;
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   303
    val _ = List.app (report_status Markup.disposed) desc;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   304
    val _ = change_commands (Graph.del_nodes desc);
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   305
  in () end;
27606
82399f3a8ca8 support for command status;
wenzelm
parents: 27600
diff changeset
   306
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   307
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   308
(* final state *)
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   309
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   310
fun the_state id =
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   311
  (case the_command id of
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   312
    {status = Finished state, ...} => state
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   313
  | {transition, ...} => error ("Unfinished command " ^ Toplevel.str_of transition));
27501
632ee56c2c0b global commands: explicit graph;
wenzelm
parents: 27438
diff changeset
   314
27428
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   315
f92d47cdc0de explicit identification of toplevel commands, with status etc.;
wenzelm
parents: 26643
diff changeset
   316
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   317
(** editor model **)
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   318
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   319
(* run commands *)
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   320
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   321
fun try_run id =
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   322
  (case try the_state (prev_command id) of
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   323
    NONE => ()
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   324
  | SOME state =>
27622
86608f598e9f editor model: run interactively for now;
wenzelm
parents: 27620
diff changeset
   325
      (case run true (#transition (the_command id)) state of
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   326
        NONE => ()
27662
34e7236443f3 tty loop: do not report status;
wenzelm
parents: 27622
diff changeset
   327
      | SOME status => report_update_status status id));
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   328
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   329
fun rerun_commands ids =
27662
34e7236443f3 tty loop: do not report status;
wenzelm
parents: 27622
diff changeset
   330
  (List.app (report_update_status Unprocessed) ids; List.app try_run ids);
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   331
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   332
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   333
(* modify document *)
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   334
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   335
fun insert_command prev id = NAMED_CRITICAL "Isar" (fn () =>
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   336
  let
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   337
    val nexts = next_commands prev;
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   338
    val _ = change_commands
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   339
     (fold (fn next => Graph.del_edge (prev, next)) nexts #> add_edge (prev, id) #>
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   340
      fold (fn next => Graph.add_edge (id, next)) nexts);
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   341
  in descendant_commands [id] end) |> rerun_commands;
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   342
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   343
fun remove_command id = NAMED_CRITICAL "Isar" (fn () =>
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   344
  let
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   345
    val prev = prev_command id;
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   346
    val nexts = next_commands id;
27620
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   347
    val _ = change_commands
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   348
     (fold (fn next => Graph.del_edge (id, next)) nexts #>
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   349
      fold (fn next => add_edge (prev, next)) nexts);
dfecac40e4be identify: more informative id in Toplevel.debug mode;
wenzelm
parents: 27616
diff changeset
   350
  in descendant_commands nexts end) |> rerun_commands;
27616
a811269b577c export type id with no_id and create_command;
wenzelm
parents: 27606
diff changeset
   351
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   352
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   353
(* concrete syntax *)
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   354
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   355
local
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   356
29348
28b0652aabd8 simplified tty model -- back to plain list history, which is independent of editor model;
wenzelm
parents: 29314
diff changeset
   357
structure P = OuterParse;
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   358
val op >> = Scan.>>;
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   359
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   360
in
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   361
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   362
val _ =
29313
6852248da4b4 Isar.command: plain Position.id;
wenzelm
parents: 28306
diff changeset
   363
  OuterSyntax.internal_command "Isar.command"
6852248da4b4 Isar.command: plain Position.id;
wenzelm
parents: 28306
diff changeset
   364
    (P.string -- P.string >> (fn (id, text) =>
6852248da4b4 Isar.command: plain Position.id;
wenzelm
parents: 28306
diff changeset
   365
      Toplevel.imperative (fn () =>
6852248da4b4 Isar.command: plain Position.id;
wenzelm
parents: 28306
diff changeset
   366
        ignore (create_command (OuterSyntax.prepare_command (Position.id id) text)))));
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   367
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   368
val _ =
29313
6852248da4b4 Isar.command: plain Position.id;
wenzelm
parents: 28306
diff changeset
   369
  OuterSyntax.internal_command "Isar.insert"
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   370
    (P.string -- P.string >> (fn (prev, id) =>
29313
6852248da4b4 Isar.command: plain Position.id;
wenzelm
parents: 28306
diff changeset
   371
      Toplevel.imperative (fn () => insert_command prev id)));
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   372
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   373
val _ =
29313
6852248da4b4 Isar.command: plain Position.id;
wenzelm
parents: 28306
diff changeset
   374
  OuterSyntax.internal_command "Isar.remove"
6852248da4b4 Isar.command: plain Position.id;
wenzelm
parents: 28306
diff changeset
   375
    (P.string >> (fn id => Toplevel.imperative (fn () => remove_command id)));
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   376
26605
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   377
end;
24e60e823d22 The global Isabelle/Isar state and main read-eval-print loop.
wenzelm
parents:
diff changeset
   378
28300
111fc1879250 moved Isar editor commands from isar_syn.ML to isar.ML;
wenzelm
parents: 27662
diff changeset
   379
end;