| author | wenzelm | 
| Fri, 05 May 2006 21:59:43 +0200 | |
| changeset 19575 | 2d9940cd52d3 | 
| parent 19545 | 98d82187392d | 
| child 19841 | f2fa72c13186 | 
| permissions | -rw-r--r-- | 
| 5828 | 1 | (* Title: Pure/Isar/toplevel.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 5 | The Isabelle/Isar toplevel. | |
| 6 | *) | |
| 7 | ||
| 8 | signature TOPLEVEL = | |
| 9 | sig | |
| 19063 | 10 | exception UNDEF | 
| 18589 | 11 | type node | 
| 12 | val theory_node: node -> theory option | |
| 13 | val proof_node: node -> ProofHistory.T option | |
| 19063 | 14 | val cases_node: (theory -> 'a) -> (Proof.state -> 'a) -> node -> 'a | 
| 15 | val body_context_node: node option -> xstring option -> Proof.context | |
| 5828 | 16 | type state | 
| 17 | val toplevel: state | |
| 7732 | 18 | val is_toplevel: state -> bool | 
| 18589 | 19 | val is_theory: state -> bool | 
| 20 | val is_proof: state -> bool | |
| 17076 | 21 | val level: state -> int | 
| 17266 | 22 | val assert: bool -> unit | 
| 6689 | 23 | val node_history_of: state -> node History.T | 
| 5828 | 24 | val node_of: state -> node | 
| 6664 | 25 | val node_case: (theory -> 'a) -> (Proof.state -> 'a) -> state -> 'a | 
| 18641 | 26 | val context_of: state -> Context.generic | 
| 5828 | 27 | val theory_of: state -> theory | 
| 28 | val proof_of: state -> Proof.state | |
| 18589 | 29 | val proof_position_of: state -> int | 
| 9453 | 30 | val enter_forward_proof: state -> Proof.state | 
| 16815 | 31 | val prompt_state_default: state -> string | 
| 32 | val prompt_state_fn: (state -> string) ref | |
| 33 | val print_state_context: state -> unit | |
| 34 | val print_state_default: bool -> state -> unit | |
| 35 | val print_state_hook: (bool -> state -> unit) -> unit | |
| 36 | val print_state_fn: (bool -> state -> unit) ref | |
| 37 | val print_state: bool -> state -> unit | |
| 38 | val pretty_state: bool -> state -> Pretty.T list | |
| 16682 | 39 | val quiet: bool ref | 
| 40 | val debug: bool ref | |
| 17321 
227f1f5c3959
added interact flag to control mode of excursions;
 wenzelm parents: 
17320diff
changeset | 41 | val interact: bool ref | 
| 16682 | 42 | val timing: bool ref | 
| 43 | val profiling: int ref | |
| 16815 | 44 | val skip_proofs: bool ref | 
| 5828 | 45 | exception TERMINATE | 
| 5990 | 46 | exception RESTART | 
| 18589 | 47 | val checkpoint: state -> state | 
| 48 | val copy: state -> state | |
| 16682 | 49 | type transition | 
| 6689 | 50 | val undo_limit: bool -> int option | 
| 5828 | 51 | val empty: transition | 
| 14923 | 52 | val name_of: transition -> string | 
| 53 | val source_of: transition -> OuterLex.token list option | |
| 5828 | 54 | val name: string -> transition -> transition | 
| 55 | val position: Position.T -> transition -> transition | |
| 14923 | 56 | val source: OuterLex.token list -> transition -> transition | 
| 5828 | 57 | val interactive: bool -> transition -> transition | 
| 58 | val print: transition -> transition | |
| 16607 | 59 | val print': string -> transition -> transition | 
| 17363 | 60 | val three_buffersN: string | 
| 61 | val print3: transition -> transition | |
| 9010 | 62 | val no_timing: transition -> transition | 
| 5828 | 63 | val reset: transition -> transition | 
| 8465 | 64 | val init: (bool -> node) -> (node -> unit) -> (node -> unit) -> transition -> transition | 
| 6689 | 65 | val exit: transition -> transition | 
| 66 | val kill: transition -> transition | |
| 5828 | 67 | val keep: (state -> unit) -> transition -> transition | 
| 7612 | 68 | val keep': (bool -> state -> unit) -> transition -> transition | 
| 6689 | 69 | val history: (node History.T -> node History.T) -> transition -> transition | 
| 5828 | 70 | val imperative: (unit -> unit) -> transition -> transition | 
| 7105 | 71 | val init_theory: (bool -> theory) -> (theory -> unit) -> (theory -> unit) | 
| 6689 | 72 | -> transition -> transition | 
| 5828 | 73 | val theory: (theory -> theory) -> transition -> transition | 
| 7612 | 74 | val theory': (bool -> theory -> theory) -> transition -> transition | 
| 18811 | 75 | val theory_context: (theory -> Proof.context * theory) -> transition -> transition | 
| 19063 | 76 | val local_theory: xstring option -> (local_theory -> local_theory) -> | 
| 77 | transition -> transition | |
| 17363 | 78 | val theory_to_proof: (theory -> Proof.state) -> transition -> transition | 
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 79 | val proof: (Proof.state -> Proof.state) -> transition -> transition | 
| 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 80 | val proofs: (Proof.state -> Proof.state Seq.seq) -> transition -> transition | 
| 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 81 | val proof': (bool -> Proof.state -> Proof.state) -> transition -> transition | 
| 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 82 | val proofs': (bool -> Proof.state -> Proof.state Seq.seq) -> transition -> transition | 
| 16815 | 83 | val actual_proof: (ProofHistory.T -> ProofHistory.T) -> transition -> transition | 
| 84 | val skip_proof: (int History.T -> int History.T) -> transition -> transition | |
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 85 | val proof_to_theory: (Proof.state -> theory) -> transition -> transition | 
| 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 86 | val proof_to_theory': (bool -> Proof.state -> theory) -> transition -> transition | 
| 18811 | 87 | val proof_to_theory_context: (bool -> Proof.state -> Proof.context * theory) | 
| 17076 | 88 | -> transition -> transition | 
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 89 | val skip_proof_to_theory: (int -> bool) -> transition -> transition | 
| 18563 | 90 | val forget_proof: transition -> transition | 
| 19063 | 91 | val present_local_theory: xstring option -> (bool -> node -> unit) -> transition -> transition | 
| 92 | val present_proof: (bool -> node -> unit) -> transition -> transition | |
| 9512 | 93 | val unknown_theory: transition -> transition | 
| 94 | val unknown_proof: transition -> transition | |
| 95 | val unknown_context: transition -> transition | |
| 5922 | 96 | val exn_message: exn -> string | 
| 5828 | 97 | val apply: bool -> transition -> state -> (state * (exn * string) option) option | 
| 17076 | 98 | val present_excursion: (transition * (state -> state -> 'a -> 'a)) list -> 'a -> 'a | 
| 5828 | 99 | val excursion: transition list -> unit | 
| 18685 | 100 | val program: (unit -> 'a) -> 'a | 
| 5828 | 101 | val set_state: state -> unit | 
| 102 | val get_state: unit -> state | |
| 103 | val exn: unit -> (exn * string) option | |
| 104 | val >> : transition -> bool | |
| 14985 | 105 | val >>> : transition list -> unit | 
| 14091 | 106 | type 'a isar | 
| 107 | val loop: 'a isar -> unit | |
| 5828 | 108 | end; | 
| 109 | ||
| 6965 | 110 | structure Toplevel: TOPLEVEL = | 
| 5828 | 111 | struct | 
| 112 | ||
| 113 | ||
| 114 | (** toplevel state **) | |
| 115 | ||
| 19063 | 116 | exception UNDEF; | 
| 117 | ||
| 118 | ||
| 16815 | 119 | (* datatype state *) | 
| 5828 | 120 | |
| 121 | datatype node = | |
| 18563 | 122 | Theory of theory * Proof.context option | (*theory with optional body context*) | 
| 123 | Proof of ProofHistory.T * theory | (*history of proof states, original theory*) | |
| 124 | SkipProof of (int History.T * theory) * theory; | |
| 125 | (*history of proof depths, resulting theory, original theory*) | |
| 5828 | 126 | |
| 18589 | 127 | val theory_node = fn Theory (thy, _) => SOME thy | _ => NONE; | 
| 128 | val proof_node = fn Proof (prf, _) => SOME prf | _ => NONE; | |
| 129 | ||
| 19063 | 130 | fun cases_node f _ (Theory (thy, _)) = f thy | 
| 131 | | cases_node _ g (Proof (prf, _)) = g (ProofHistory.current prf) | |
| 132 | | cases_node f _ (SkipProof ((_, thy), _)) = f thy; | |
| 133 | ||
| 134 | fun body_context_node (SOME (Theory (_, SOME ctxt))) NONE = ctxt | |
| 135 | | body_context_node (SOME node) loc = | |
| 136 | node |> cases_node (LocalTheory.init loc) | |
| 137 | (if is_some loc then LocalTheory.init loc o Proof.theory_of | |
| 138 | else Proof.context_of) | |
| 139 | | body_context_node NONE _ = raise UNDEF; | |
| 140 | ||
| 8465 | 141 | datatype state = State of (node History.T * ((node -> unit) * (node -> unit))) option; | 
| 5828 | 142 | |
| 15531 | 143 | val toplevel = State NONE; | 
| 5828 | 144 | |
| 15531 | 145 | fun is_toplevel (State NONE) = true | 
| 7732 | 146 | | is_toplevel _ = false; | 
| 147 | ||
| 17076 | 148 | fun level (State NONE) = 0 | 
| 149 | | level (State (SOME (node, _))) = | |
| 150 | (case History.current node of | |
| 151 | Theory _ => 0 | |
| 18563 | 152 | | Proof (prf, _) => Proof.level (ProofHistory.current prf) | 
| 153 | | SkipProof ((h, _), _) => History.current h + 1); (*different notion of proof depth!*) | |
| 17076 | 154 | |
| 15531 | 155 | fun str_of_state (State NONE) = "at top level" | 
| 16815 | 156 | | str_of_state (State (SOME (node, _))) = | 
| 157 | (case History.current node of | |
| 158 | Theory _ => "in theory mode" | |
| 159 | | Proof _ => "in proof mode" | |
| 160 | | SkipProof _ => "in skipped proof mode"); | |
| 5946 
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
 wenzelm parents: 
5939diff
changeset | 161 | |
| 
a4600d21b59b
print_state hook, obeys Goals.current_goals_markers by default;
 wenzelm parents: 
5939diff
changeset | 162 | |
| 5828 | 163 | (* top node *) | 
| 164 | ||
| 17266 | 165 | fun assert true = () | 
| 166 | | assert false = raise UNDEF; | |
| 167 | ||
| 15531 | 168 | fun node_history_of (State NONE) = raise UNDEF | 
| 169 | | node_history_of (State (SOME (node, _))) = node; | |
| 6689 | 170 | |
| 171 | val node_of = History.current o node_history_of; | |
| 5828 | 172 | |
| 18589 | 173 | fun is_theory state = not (is_toplevel state) andalso is_some (theory_node (node_of state)); | 
| 174 | fun is_proof state = not (is_toplevel state) andalso is_some (proof_node (node_of state)); | |
| 175 | ||
| 19063 | 176 | fun node_case f g state = cases_node f g (node_of state); | 
| 5828 | 177 | |
| 18641 | 178 | val context_of = node_case Context.Theory (Context.Proof o Proof.context_of); | 
| 6664 | 179 | val theory_of = node_case I Proof.theory_of; | 
| 18589 | 180 | val proof_of = node_case (fn _ => raise UNDEF) I; | 
| 17208 | 181 | |
| 18589 | 182 | fun proof_position_of state = | 
| 183 | (case node_of state of | |
| 184 | Proof (prf, _) => ProofHistory.position prf | |
| 185 | | _ => raise UNDEF); | |
| 6664 | 186 | |
| 17363 | 187 | val enter_forward_proof = node_case (Proof.init o ProofContext.init) Proof.enter_forward; | 
| 5828 | 188 | |
| 189 | ||
| 16815 | 190 | (* prompt state *) | 
| 191 | ||
| 192 | fun prompt_state_default (State _) = Source.default_prompt; | |
| 193 | ||
| 194 | val prompt_state_fn = ref prompt_state_default; | |
| 195 | fun prompt_state state = ! prompt_state_fn state; | |
| 196 | ||
| 197 | ||
| 198 | (* print state *) | |
| 199 | ||
| 200 | fun pretty_context thy = [Pretty.block | |
| 201 | [Pretty.str "theory", Pretty.brk 1, Pretty.str (Context.theory_name thy), | |
| 202 | Pretty.str " =", Pretty.brk 1, ThyInfo.pretty_theory thy]]; | |
| 203 | ||
| 204 | fun pretty_state_context state = | |
| 205 | (case try theory_of state of NONE => [] | |
| 206 | | SOME thy => pretty_context thy); | |
| 207 | ||
| 17076 | 208 | fun pretty_node prf_only (Theory (thy, _)) = if prf_only then [] else pretty_context thy | 
| 18563 | 209 | | pretty_node _ (Proof (prf, _)) = | 
| 16815 | 210 | Proof.pretty_state (ProofHistory.position prf) (ProofHistory.current prf) | 
| 18563 | 211 | | pretty_node _ (SkipProof ((h, _), _)) = | 
| 16815 | 212 |       [Pretty.str ("skipped proof: depth " ^ string_of_int (History.current h))];
 | 
| 213 | ||
| 214 | fun pretty_state prf_only state = | |
| 215 | let val ref (begin_state, end_state, _) = Display.current_goals_markers in | |
| 216 | (case try node_of state of NONE => [] | |
| 217 | | SOME node => | |
| 218 | (if begin_state = "" then [] else [Pretty.str begin_state]) @ | |
| 219 | pretty_node prf_only node @ | |
| 220 | (if end_state = "" then [] else [Pretty.str end_state])) | |
| 221 | end; | |
| 222 | ||
| 223 | val print_state_context = Pretty.writelns o pretty_state_context; | |
| 224 | fun print_state_default prf_only state = Pretty.writelns (pretty_state prf_only state); | |
| 225 | ||
| 226 | val print_state_hooks = ref ([]: (bool -> state -> unit) list); | |
| 227 | fun print_state_hook f = change print_state_hooks (cons f); | |
| 228 | val print_state_fn = ref print_state_default; | |
| 229 | ||
| 230 | fun print_state prf_only state = | |
| 231 | (List.app (fn f => (try (fn () => f prf_only state) (); ())) (! print_state_hooks); | |
| 232 | ! print_state_fn prf_only state); | |
| 233 | ||
| 234 | ||
| 15668 | 235 | |
| 5828 | 236 | (** toplevel transitions **) | 
| 237 | ||
| 16682 | 238 | val quiet = ref false; | 
| 239 | val debug = ref false; | |
| 17321 
227f1f5c3959
added interact flag to control mode of excursions;
 wenzelm parents: 
17320diff
changeset | 240 | val interact = ref false; | 
| 16682 | 241 | val timing = Output.timing; | 
| 242 | val profiling = ref 0; | |
| 16815 | 243 | val skip_proofs = ref false; | 
| 16682 | 244 | |
| 5828 | 245 | exception TERMINATE; | 
| 5990 | 246 | exception RESTART; | 
| 7022 | 247 | exception EXCURSION_FAIL of exn * string; | 
| 6689 | 248 | exception FAILURE of state * exn; | 
| 5828 | 249 | |
| 18685 | 250 | fun debugging f x = | 
| 251 | if ! debug then | |
| 252 | setmp Library.do_transform_failure false | |
| 253 | exception_trace (fn () => f x) | |
| 254 | else f x; | |
| 255 | ||
| 5828 | 256 | |
| 16815 | 257 | (* node transactions and recovery from stale theories *) | 
| 6689 | 258 | |
| 16815 | 259 | (*NB: proof commands should be non-destructive!*) | 
| 7022 | 260 | |
| 6689 | 261 | local | 
| 262 | ||
| 16452 | 263 | fun is_stale state = Context.is_stale (theory_of state) handle UNDEF => false; | 
| 6689 | 264 | |
| 18685 | 265 | val stale_theory = ERROR "Stale theory encountered after succesful execution!"; | 
| 16815 | 266 | |
| 19101 | 267 | fun checkpoint_node (Theory (thy, _)) = Theory (Theory.checkpoint thy, NONE) | 
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 268 | | checkpoint_node node = node; | 
| 6689 | 269 | |
| 19101 | 270 | fun copy_node (Theory (thy, _)) = Theory (Theory.copy thy, NONE) | 
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 271 | | copy_node node = node; | 
| 6689 | 272 | |
| 15531 | 273 | fun return (result, NONE) = result | 
| 274 | | return (result, SOME exn) = raise FAILURE (result, exn); | |
| 7022 | 275 | |
| 16729 | 276 | fun interruptible f x = | 
| 277 | let val y = ref x | |
| 278 | in raise_interrupt (fn () => y := f x) (); ! y end; | |
| 279 | ||
| 6689 | 280 | in | 
| 281 | ||
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 282 | fun transaction _ _ (State NONE) = raise UNDEF | 
| 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 283 | | transaction hist f (State (SOME (node, term))) = | 
| 6689 | 284 | let | 
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 285 | val cont_node = History.map checkpoint_node node; | 
| 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 286 | val back_node = History.map copy_node cont_node; | 
| 16729 | 287 | fun state nd = State (SOME (nd, term)); | 
| 288 | fun normal_state nd = (state nd, NONE); | |
| 289 | fun error_state nd exn = (state nd, SOME exn); | |
| 6689 | 290 | |
| 16729 | 291 | val (result, err) = | 
| 292 | cont_node | |
| 18563 | 293 | |> (f | 
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 294 | |> (if hist then History.apply_copy copy_node else History.map) | 
| 18685 | 295 | |> debugging | 
| 16729 | 296 | |> interruptible | 
| 18685 | 297 | |> setmp Output.do_toplevel_errors false) | 
| 16729 | 298 | |> normal_state | 
| 299 | handle exn => error_state cont_node exn; | |
| 6689 | 300 | in | 
| 16729 | 301 | if is_stale result | 
| 18955 | 302 | then return (error_state back_node (the_default stale_theory err)) | 
| 16729 | 303 | else return (result, err) | 
| 6689 | 304 | end; | 
| 305 | ||
| 18589 | 306 | fun mapping f (State (SOME (node, term))) = State (SOME (History.map f node, term)) | 
| 307 | | mapping _ state = state; | |
| 308 | ||
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 309 | val checkpoint = mapping checkpoint_node; | 
| 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 310 | val copy = mapping copy_node; | 
| 18589 | 311 | |
| 6689 | 312 | end; | 
| 313 | ||
| 314 | ||
| 315 | (* primitive transitions *) | |
| 316 | ||
| 18563 | 317 | (*Note: Recovery from stale theories is provided only for theory-level | 
| 18589 | 318 | operations via Transaction. Other node or state operations should | 
| 319 | not touch theories at all. Interrupts are enabled only for Keep and | |
| 320 | Transaction.*) | |
| 5828 | 321 | |
| 322 | datatype trans = | |
| 18563 | 323 | Reset | (*empty toplevel*) | 
| 8465 | 324 | Init of (bool -> node) * ((node -> unit) * (node -> unit)) | | 
| 18563 | 325 | (*init node; with exit/kill operation*) | 
| 326 | Exit | (*conclude node*) | |
| 327 | Kill | (*abort node*) | |
| 328 | Keep of bool -> state -> unit | (*peek at state*) | |
| 329 | History of node History.T -> node History.T | (*history operation (undo etc.)*) | |
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 330 | Transaction of bool * (bool -> node -> node); (*node transaction*) | 
| 6689 | 331 | |
| 15531 | 332 | fun undo_limit int = if int then NONE else SOME 0; | 
| 6689 | 333 | |
| 334 | local | |
| 5828 | 335 | |
| 6689 | 336 | fun apply_tr _ Reset _ = toplevel | 
| 15531 | 337 | | apply_tr int (Init (f, term)) (State NONE) = | 
| 338 | State (SOME (History.init (undo_limit int) (f int), term)) | |
| 339 | | apply_tr _ (Init _ ) (State (SOME _)) = raise UNDEF | |
| 340 | | apply_tr _ Exit (State NONE) = raise UNDEF | |
| 341 | | apply_tr _ Exit (State (SOME (node, (exit, _)))) = | |
| 342 | (exit (History.current node); State NONE) | |
| 343 | | apply_tr _ Kill (State NONE) = raise UNDEF | |
| 344 | | apply_tr _ Kill (State (SOME (node, (_, kill)))) = | |
| 345 | (kill (History.current node); State NONE) | |
| 18718 | 346 | | apply_tr int (Keep f) state = | 
| 347 | (setmp Output.do_toplevel_errors false (raise_interrupt (f int)) state; state) | |
| 15531 | 348 | | apply_tr _ (History _) (State NONE) = raise UNDEF | 
| 349 | | apply_tr _ (History f) (State (SOME (node, term))) = State (SOME (f node, term)) | |
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 350 | | apply_tr int (Transaction (hist, f)) state = transaction hist (fn x => f int x) state; | 
| 5828 | 351 | |
| 6689 | 352 | fun apply_union _ [] state = raise FAILURE (state, UNDEF) | 
| 353 | | apply_union int (tr :: trs) state = | |
| 18685 | 354 | apply_tr int tr state | 
| 6689 | 355 | handle UNDEF => apply_union int trs state | 
| 356 | | FAILURE (alt_state, UNDEF) => apply_union int trs alt_state | |
| 357 | | exn as FAILURE _ => raise exn | |
| 358 | | exn => raise FAILURE (state, exn); | |
| 359 | ||
| 360 | in | |
| 361 | ||
| 15531 | 362 | fun apply_trans int trs state = (apply_union int trs state, NONE) | 
| 363 | handle FAILURE (alt_state, exn) => (alt_state, SOME exn) | exn => (state, SOME exn); | |
| 6689 | 364 | |
| 365 | end; | |
| 5828 | 366 | |
| 367 | ||
| 368 | (* datatype transition *) | |
| 369 | ||
| 370 | datatype transition = Transition of | |
| 16815 | 371 |  {name: string,                        (*command name*)
 | 
| 372 | pos: Position.T, (*source position*) | |
| 373 | source: OuterLex.token list option, (*source text*) | |
| 374 | int_only: bool, (*interactive-only*) | |
| 375 | print: string list, (*print modes (union)*) | |
| 376 | no_timing: bool, (*suppress timing*) | |
| 377 | trans: trans list}; (*primitive transitions (union)*) | |
| 5828 | 378 | |
| 14923 | 379 | fun make_transition (name, pos, source, int_only, print, no_timing, trans) = | 
| 380 |   Transition {name = name, pos = pos, source = source,
 | |
| 381 | int_only = int_only, print = print, no_timing = no_timing, trans = trans}; | |
| 5828 | 382 | |
| 14923 | 383 | fun map_transition f (Transition {name, pos, source, int_only, print, no_timing, trans}) =
 | 
| 384 | make_transition (f (name, pos, source, int_only, print, no_timing, trans)); | |
| 5828 | 385 | |
| 16607 | 386 | val empty = make_transition ("<unknown>", Position.none, NONE, false, [], false, []);
 | 
| 14923 | 387 | |
| 388 | fun name_of (Transition {name, ...}) = name;
 | |
| 389 | fun source_of (Transition {source, ...}) = source;
 | |
| 5828 | 390 | |
| 391 | ||
| 392 | (* diagnostics *) | |
| 393 | ||
| 394 | fun str_of_transition (Transition {name, pos, ...}) = quote name ^ Position.str_of pos;
 | |
| 395 | ||
| 396 | fun command_msg msg tr = msg ^ "command " ^ str_of_transition tr; | |
| 397 | fun at_command tr = command_msg "At " tr ^ "."; | |
| 398 | ||
| 399 | fun type_error tr state = | |
| 18685 | 400 | ERROR (command_msg "Illegal application of " tr ^ " " ^ str_of_state state); | 
| 5828 | 401 | |
| 402 | ||
| 403 | (* modify transitions *) | |
| 404 | ||
| 14923 | 405 | fun name nm = map_transition (fn (_, pos, source, int_only, print, no_timing, trans) => | 
| 406 | (nm, pos, source, int_only, print, no_timing, trans)); | |
| 5828 | 407 | |
| 14923 | 408 | fun position pos = map_transition (fn (name, _, source, int_only, print, no_timing, trans) => | 
| 409 | (name, pos, source, int_only, print, no_timing, trans)); | |
| 9010 | 410 | |
| 14923 | 411 | fun source src = map_transition (fn (name, pos, _, int_only, print, no_timing, trans) => | 
| 15531 | 412 | (name, pos, SOME src, int_only, print, no_timing, trans)); | 
| 5828 | 413 | |
| 14923 | 414 | fun interactive int_only = map_transition (fn (name, pos, source, _, print, no_timing, trans) => | 
| 415 | (name, pos, source, int_only, print, no_timing, trans)); | |
| 416 | ||
| 17363 | 417 | val no_timing = map_transition (fn (name, pos, source, int_only, print, _, trans) => | 
| 418 | (name, pos, source, int_only, print, true, trans)); | |
| 419 | ||
| 420 | fun add_trans tr = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) => | |
| 421 | (name, pos, source, int_only, print, no_timing, trans @ [tr])); | |
| 422 | ||
| 16607 | 423 | fun print' mode = map_transition (fn (name, pos, source, int_only, print, no_timing, trans) => | 
| 424 | (name, pos, source, int_only, insert (op =) mode print, no_timing, trans)); | |
| 425 | ||
| 426 | val print = print' ""; | |
| 5828 | 427 | |
| 17363 | 428 | val three_buffersN = "three_buffers"; | 
| 429 | val print3 = print' three_buffersN; | |
| 5828 | 430 | |
| 431 | ||
| 432 | (* build transitions *) | |
| 433 | ||
| 434 | val reset = add_trans Reset; | |
| 6689 | 435 | fun init f exit kill = add_trans (Init (f, (exit, kill))); | 
| 436 | val exit = add_trans Exit; | |
| 437 | val kill = add_trans Kill; | |
| 7612 | 438 | val keep' = add_trans o Keep; | 
| 6689 | 439 | val history = add_trans o History; | 
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 440 | fun map_current f = add_trans (Transaction (false, f)); | 
| 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 441 | fun app_current f = add_trans (Transaction (true, f)); | 
| 5828 | 442 | |
| 7612 | 443 | fun keep f = add_trans (Keep (fn _ => f)); | 
| 5828 | 444 | fun imperative f = keep (fn _ => f ()); | 
| 445 | ||
| 6689 | 446 | fun init_theory f exit kill = | 
| 17076 | 447 | init (fn int => Theory (f int, NONE)) | 
| 448 | (fn Theory (thy, _) => exit thy | _ => raise UNDEF) | |
| 449 | (fn Theory (thy, _) => kill thy | _ => raise UNDEF); | |
| 5828 | 450 | |
| 15668 | 451 | |
| 452 | (* typed transitions *) | |
| 453 | ||
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 454 | fun theory f = app_current | 
| 18563 | 455 | (K (fn Theory (thy, _) => Theory (f thy, NONE) | _ => raise UNDEF)); | 
| 456 | ||
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 457 | fun theory' f = app_current (fn int => | 
| 17076 | 458 | (fn Theory (thy, _) => Theory (f int thy, NONE) | _ => raise UNDEF)); | 
| 459 | ||
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 460 | fun theory_context f = app_current | 
| 18811 | 461 | (K (fn Theory (thy, _) => Theory (swap (apfst SOME (f thy))) | _ => raise UNDEF)); | 
| 15668 | 462 | |
| 19063 | 463 | fun local_theory loc f = theory_context (LocalTheory.mapping loc f); | 
| 18955 | 464 | |
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 465 | fun theory_to_proof f = app_current (fn int => | 
| 17076 | 466 | (fn Theory (thy, _) => | 
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 467 | if ! skip_proofs then | 
| 18811 | 468 | SkipProof ((History.init (undo_limit int) 0, #2 (Proof.global_skip_proof int (f thy))), thy) | 
| 18592 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 469 | else Proof (ProofHistory.init (undo_limit int) (f thy), thy) | 
| 
451d622bb4a9
transactions now always see quasi-functional intermediate checkpoint;
 wenzelm parents: 
18589diff
changeset | 470 | | _ => raise UNDEF)); | 
| 18563 | 471 | |
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 472 | fun proofs' f = map_current (fn int => | 
| 18563 | 473 | (fn Proof (prf, orig_thy) => Proof (ProofHistory.applys (f int) prf, orig_thy) | 
| 474 | | SkipProof ((h, thy), orig_thy) => SkipProof ((History.apply I h, thy), orig_thy) | |
| 16815 | 475 | | _ => raise UNDEF)); | 
| 15668 | 476 | |
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 477 | fun proof' f = proofs' (Seq.single oo f); | 
| 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 478 | val proofs = proofs' o K; | 
| 6689 | 479 | val proof = proof' o K; | 
| 16815 | 480 | |
| 481 | fun actual_proof f = map_current (fn _ => | |
| 18563 | 482 | (fn Proof (prf, orig_thy) => Proof (f prf, orig_thy) | _ => raise UNDEF)); | 
| 16815 | 483 | |
| 484 | fun skip_proof f = map_current (fn _ => | |
| 18563 | 485 | (fn SkipProof ((h, thy), orig_thy) => SkipProof ((f h, thy), orig_thy) | _ => raise UNDEF)); | 
| 15668 | 486 | |
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 487 | fun end_proof f = map_current (fn int => | 
| 18563 | 488 | (fn Proof (prf, _) => Theory (f int (ProofHistory.current prf)) | 
| 489 | | SkipProof ((h, thy), _) => if History.current h = 0 then Theory (thy, NONE) else raise UNDEF | |
| 490 | | _ => raise UNDEF)); | |
| 491 | ||
| 492 | val forget_proof = map_current (fn _ => | |
| 493 | (fn Proof (_, orig_thy) => Theory (orig_thy, NONE) | |
| 494 | | SkipProof (_, orig_thy) => Theory (orig_thy, NONE) | |
| 15237 
250e9be7a09d
Some changes to allow skipping of proof scripts.
 berghofe parents: 
14985diff
changeset | 495 | | _ => raise UNDEF)); | 
| 15668 | 496 | |
| 17076 | 497 | fun proof_to_theory' f = end_proof (rpair NONE oo f); | 
| 498 | fun proof_to_theory f = proof_to_theory' (K f); | |
| 18811 | 499 | fun proof_to_theory_context f = end_proof ((swap o apfst SOME) oo f); | 
| 15668 | 500 | |
| 16815 | 501 | fun skip_proof_to_theory p = map_current (fn _ => | 
| 18563 | 502 | (fn SkipProof ((h, thy), _) => | 
| 17904 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 503 | if p (History.current h) then Theory (thy, NONE) | 
| 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 504 | else raise UNDEF | 
| 
21c6894b5998
simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
 wenzelm parents: 
17513diff
changeset | 505 | | _ => raise UNDEF)); | 
| 5828 | 506 | |
| 19063 | 507 | fun present_local_theory loc f = app_current (fn int => | 
| 508 | (fn Theory (thy, _) => Theory (swap (apfst SOME (LocalTheory.mapping loc I thy))) | |
| 509 | | _ => raise UNDEF) #> tap (f int)); | |
| 510 | ||
| 511 | fun present_proof f = map_current (fn int => | |
| 512 | (fn node as Proof _ => node | _ => raise UNDEF) #> tap (f int)); | |
| 513 | ||
| 9512 | 514 | val unknown_theory = imperative (fn () => warning "Unknown theory context"); | 
| 515 | val unknown_proof = imperative (fn () => warning "Unknown proof context"); | |
| 516 | val unknown_context = imperative (fn () => warning "Unknown context"); | |
| 517 | ||
| 5828 | 518 | |
| 519 | ||
| 520 | (** toplevel transactions **) | |
| 521 | ||
| 522 | (* print exceptions *) | |
| 523 | ||
| 15668 | 524 | local | 
| 525 | ||
| 526 | fun with_context f xs = | |
| 527 | (case Context.get_context () of NONE => [] | |
| 16452 | 528 | | SOME thy => map (f thy) xs); | 
| 15668 | 529 | |
| 530 | fun raised name [] = "exception " ^ name ^ " raised" | |
| 531 | | raised name [msg] = "exception " ^ name ^ " raised: " ^ msg | |
| 532 |   | raised name msgs = cat_lines (("exception " ^ name ^ " raised:") :: msgs);
 | |
| 5828 | 533 | |
| 15668 | 534 | fun exn_msg _ TERMINATE = "Exit." | 
| 535 | | exn_msg _ RESTART = "Restart." | |
| 536 | | exn_msg _ Interrupt = "Interrupt." | |
| 18685 | 537 | | exn_msg _ Output.TOPLEVEL_ERROR = "Error." | 
| 19545 | 538 | | exn_msg _ (SYS_ERROR msg) = "## SYSTEM ERROR ##\n" ^ msg | 
| 18685 | 539 | | exn_msg _ (ERROR msg) = msg | 
| 17363 | 540 | | exn_msg detailed (EXCEPTION (exn, msg)) = cat_lines [exn_msg detailed exn, msg] | 
| 16452 | 541 | | exn_msg detailed (EXCURSION_FAIL (exn, msg)) = cat_lines [exn_msg detailed exn, msg] | 
| 15668 | 542 | | exn_msg false (THEORY (msg, _)) = msg | 
| 16452 | 543 | | exn_msg true (THEORY (msg, thys)) = raised "THEORY" (msg :: map Context.str_of_thy thys) | 
| 15668 | 544 | | exn_msg detailed (MetaSimplifier.SIMPROC_FAIL (name, exn)) = | 
| 545 | fail_msg detailed "simproc" ((name, Position.none), exn) | |
| 546 | | exn_msg detailed (Attrib.ATTRIB_FAIL info) = fail_msg detailed "attribute" info | |
| 547 | | exn_msg detailed (Method.METHOD_FAIL info) = fail_msg detailed "method" info | |
| 548 | | exn_msg detailed (Antiquote.ANTIQUOTE_FAIL info) = fail_msg detailed "antiquotation" info | |
| 549 | | exn_msg false (Syntax.AST (msg, _)) = raised "AST" [msg] | |
| 550 | | exn_msg true (Syntax.AST (msg, asts)) = | |
| 551 | raised "AST" (msg :: map (Pretty.string_of o Syntax.pretty_ast) asts) | |
| 552 | | exn_msg false (TYPE (msg, _, _)) = raised "TYPE" [msg] | |
| 553 | | exn_msg true (TYPE (msg, Ts, ts)) = raised "TYPE" (msg :: | |
| 554 | with_context Sign.string_of_typ Ts @ with_context Sign.string_of_term ts) | |
| 555 | | exn_msg false (TERM (msg, _)) = raised "TERM" [msg] | |
| 556 | | exn_msg true (TERM (msg, ts)) = raised "TERM" (msg :: with_context Sign.string_of_term ts) | |
| 557 | | exn_msg false (THM (msg, _, _)) = raised "THM" [msg] | |
| 558 | | exn_msg true (THM (msg, i, thms)) = | |
| 559 |       raised ("THM " ^ string_of_int i) (msg :: map Display.string_of_thm thms)
 | |
| 560 | | exn_msg _ Option = raised "Option" [] | |
| 561 | | exn_msg _ UnequalLengths = raised "UnequalLengths" [] | |
| 562 | | exn_msg _ Empty = raised "Empty" [] | |
| 563 | | exn_msg _ Subscript = raised "Subscript" [] | |
| 16197 | 564 | | exn_msg _ (Fail msg) = raised "Fail" [msg] | 
| 15668 | 565 | | exn_msg _ exn = General.exnMessage exn | 
| 566 | and fail_msg detailed kind ((name, pos), exn) = | |
| 567 | "Error in " ^ kind ^ " " ^ quote name ^ Position.str_of pos ^ ":\n" ^ exn_msg detailed exn; | |
| 15519 | 568 | |
| 15668 | 569 | in | 
| 570 | ||
| 571 | fun exn_message exn = exn_msg (! debug) exn; | |
| 5828 | 572 | |
| 15531 | 573 | fun print_exn NONE = () | 
| 18685 | 574 | | print_exn (SOME (exn, s)) = Output.error_msg (cat_lines [exn_message exn, s]); | 
| 5828 | 575 | |
| 15668 | 576 | end; | 
| 577 | ||
| 5828 | 578 | |
| 579 | (* apply transitions *) | |
| 580 | ||
| 6664 | 581 | local | 
| 582 | ||
| 9010 | 583 | fun app int (tr as Transition {trans, int_only, print, no_timing, ...}) state =
 | 
| 5828 | 584 | let | 
| 16682 | 585 | val _ = conditional (not int andalso int_only) (fn () => | 
| 586 | warning (command_msg "Interactive-only " tr)); | |
| 587 | ||
| 18685 | 588 | fun do_timing f x = (Output.info (command_msg "" tr); timeap f x); | 
| 16682 | 589 | fun do_profiling f x = profile (! profiling) f x; | 
| 590 | ||
| 6689 | 591 | val (result, opt_exn) = | 
| 16729 | 592 | state |> (apply_trans int trans | 
| 593 | |> (if ! profiling > 0 then do_profiling else I) | |
| 594 | |> (if ! profiling > 0 orelse ! timing andalso not no_timing then do_timing else I)); | |
| 16607 | 595 | val _ = conditional (int andalso not (! quiet) andalso | 
| 596 |         exists (fn m => m mem_string print) ("" :: ! print_mode))
 | |
| 597 | (fn () => print_state false result); | |
| 15570 | 598 | in (result, Option.map (fn UNDEF => type_error tr state | exn => exn) opt_exn) end; | 
| 6664 | 599 | |
| 600 | in | |
| 5828 | 601 | |
| 6689 | 602 | fun apply int tr st = | 
| 6965 | 603 | (case app int tr st of | 
| 15531 | 604 | (_, SOME TERMINATE) => NONE | 
| 605 | | (_, SOME RESTART) => SOME (toplevel, NONE) | |
| 606 | | (state', SOME (EXCURSION_FAIL exn_info)) => SOME (state', SOME exn_info) | |
| 607 | | (state', SOME exn) => SOME (state', SOME (exn, at_command tr)) | |
| 608 | | (state', NONE) => SOME (state', NONE)); | |
| 6664 | 609 | |
| 610 | end; | |
| 5828 | 611 | |
| 612 | ||
| 17076 | 613 | (* excursion: toplevel -- apply transformers/presentation -- toplevel *) | 
| 5828 | 614 | |
| 6664 | 615 | local | 
| 616 | ||
| 5828 | 617 | fun excur [] x = x | 
| 17076 | 618 | | excur ((tr, pr) :: trs) (st, res) = | 
| 17321 
227f1f5c3959
added interact flag to control mode of excursions;
 wenzelm parents: 
17320diff
changeset | 619 | (case apply (! interact) tr st of | 
| 15531 | 620 | SOME (st', NONE) => | 
| 18685 | 621 | excur trs (st', pr st st' res handle exn => | 
| 10324 | 622 | raise EXCURSION_FAIL (exn, "Presentation failed\n" ^ at_command tr)) | 
| 15531 | 623 | | SOME (st', SOME exn_info) => raise EXCURSION_FAIL exn_info | 
| 624 | | NONE => raise EXCURSION_FAIL (TERMINATE, at_command tr)); | |
| 5828 | 625 | |
| 17076 | 626 | fun no_pr _ _ _ = (); | 
| 627 | ||
| 6664 | 628 | in | 
| 629 | ||
| 17076 | 630 | fun present_excursion trs res = | 
| 15531 | 631 | (case excur trs (State NONE, res) of | 
| 632 | (State NONE, res') => res' | |
| 18685 | 633 | | _ => error "Unfinished development at end of input") | 
| 9134 | 634 | handle exn => error (exn_message exn); | 
| 635 | ||
| 17076 | 636 | fun excursion trs = present_excursion (map (rpair no_pr) trs) (); | 
| 7062 | 637 | |
| 6664 | 638 | end; | 
| 639 | ||
| 5828 | 640 | |
| 18718 | 641 | (* toplevel program: independent of state *) | 
| 18685 | 642 | |
| 643 | fun program f = | |
| 18718 | 644 | Output.ML_errors (fn () => debugging f () handle exn => error (exn_message exn)) (); | 
| 18685 | 645 | |
| 646 | ||
| 5828 | 647 | |
| 648 | (** interactive transformations **) | |
| 649 | ||
| 650 | (* the global state reference *) | |
| 651 | ||
| 15531 | 652 | val global_state = ref (toplevel, NONE: (exn * string) option); | 
| 5828 | 653 | |
| 15531 | 654 | fun set_state state = global_state := (state, NONE); | 
| 5828 | 655 | fun get_state () = fst (! global_state); | 
| 656 | fun exn () = snd (! global_state); | |
| 657 | ||
| 658 | ||
| 6965 | 659 | (* the Isar source of transitions *) | 
| 660 | ||
| 14091 | 661 | type 'a isar = | 
| 6965 | 662 | (transition, (transition option, | 
| 12881 | 663 | (OuterLex.token, (OuterLex.token option, (OuterLex.token, (OuterLex.token, | 
| 14091 | 664 | Position.T * (Symbol.symbol, (string, 'a) Source.source) Source.source) | 
| 12881 | 665 | Source.source) Source.source) Source.source) Source.source) Source.source) Source.source; | 
| 6965 | 666 | |
| 667 | ||
| 5828 | 668 | (* apply transformers to global state *) | 
| 669 | ||
| 14985 | 670 | nonfix >> >>>; | 
| 5828 | 671 | |
| 672 | fun >> tr = | |
| 673 | (case apply true tr (get_state ()) of | |
| 15531 | 674 | NONE => false | 
| 675 | | SOME (state', exn_info) => | |
| 5828 | 676 | (global_state := (state', exn_info); | 
| 16729 | 677 | print_exn exn_info; | 
| 5828 | 678 | true)); | 
| 679 | ||
| 14985 | 680 | fun >>> [] = () | 
| 681 | | >>> (tr :: trs) = if >> tr then >>> trs else (); | |
| 682 | ||
| 17513 
0393718c2f1c
get_interrupt: special handling of IO.io now in ML-Systems/smlnj-basis-compat.ML;
 wenzelm parents: 
17500diff
changeset | 683 | (*Spurious interrupts ahead! Race condition?*) | 
| 
0393718c2f1c
get_interrupt: special handling of IO.io now in ML-Systems/smlnj-basis-compat.ML;
 wenzelm parents: 
17500diff
changeset | 684 | fun get_interrupt src = SOME (Source.get_single src) handle Interrupt => NONE; | 
| 7602 | 685 | |
| 5828 | 686 | fun raw_loop src = | 
| 7602 | 687 | (case get_interrupt (Source.set_prompt (prompt_state (get_state ())) src) of | 
| 15531 | 688 | NONE => (writeln "\nInterrupt."; raw_loop src) | 
| 689 | | SOME NONE => () | |
| 690 | | SOME (SOME (tr, src')) => if >> tr then raw_loop src' else ()); | |
| 5828 | 691 | |
| 12987 | 692 | fun loop src = ignore_interrupt raw_loop src; | 
| 5828 | 693 | |
| 694 | end; |