| author | wenzelm | 
| Sat, 29 Dec 2001 18:36:12 +0100 | |
| changeset 12610 | 8b9845807f77 | 
| parent 8807 | 0046be1769f9 | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 1 | (* Title: Pure/Isar/proof_history.ML | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 3 | Author: Markus Wenzel, TU Muenchen | 
| 8807 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 5 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 6 | Histories of proof states, with undo / redo and prev / back. | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 7 | *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 8 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 9 | signature PROOF_HISTORY = | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 10 | sig | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 11 | type T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 12 | exception FAIL of string | 
| 6902 | 13 | val position: T -> int | 
| 6684 | 14 | val init: int option -> Proof.state -> T | 
| 15 | val is_initial: T -> bool | |
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 16 | val current: T -> Proof.state | 
| 7731 | 17 | val clear: int -> T -> T | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 18 | val undo: T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 19 | val redo: T -> T | 
| 7366 | 20 | val back: bool -> T -> T | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 21 | val applys: (Proof.state -> Proof.state Seq.seq) -> T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 22 | val apply: (Proof.state -> Proof.state) -> T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 23 | end; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 24 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 25 | structure ProofHistory: PROOF_HISTORY = | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 26 | struct | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 27 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 28 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 29 | (* datatype proof history *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 30 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 31 | type node = | 
| 6902 | 32 | Proof.state * (*first result*) | 
| 33 | Proof.state Seq.seq; (*alternative results*) | |
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 34 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 35 | type proof = node * node list; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 36 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 37 | datatype T = ProofHistory of proof History.T; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 38 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 39 | exception FAIL of string; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 40 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 41 | fun app f (ProofHistory x) = ProofHistory (f x); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 42 | fun hist_app f = app (History.apply f); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 43 | |
| 6902 | 44 | fun position (ProofHistory prf) = length (snd (History.current prf)); | 
| 45 | ||
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 46 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 47 | (* generic history operations *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 48 | |
| 6902 | 49 | fun init lim st = ProofHistory (History.init lim ((st, Seq.empty), [])); | 
| 6684 | 50 | fun is_initial (ProofHistory prf) = History.is_initial prf; | 
| 6902 | 51 | fun current (ProofHistory prf) = fst (fst (History.current prf)); | 
| 7731 | 52 | val clear = app o History.clear; | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 53 | val undo = app History.undo; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 54 | val redo = app History.redo; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 55 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 56 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 57 | (* backtrack *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 58 | |
| 7366 | 59 | fun back_fun recur ((_, stq), nodes) = | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 60 | (case Seq.pull stq of | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 61 | None => | 
| 7366 | 62 | if recur andalso not (null nodes) then | 
| 63 | (writeln "back: trying previous proof state ..."; back_fun recur (hd nodes, tl nodes)) | |
| 64 | else raise FAIL "back: no alternatives" | |
| 6902 | 65 | | Some result => (result, nodes)); | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 66 | |
| 7366 | 67 | val back = hist_app o back_fun; | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 68 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 69 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 70 | (* apply transformer *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 71 | |
| 6902 | 72 | fun applys f = hist_app (fn (node as (st, _), nodes) => | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 73 | (case Seq.pull (f st) of | 
| 6902 | 74 | None => raise FAIL "empty result sequence -- proof command failed" | 
| 75 | | Some results => (results, node :: nodes))); | |
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 76 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 77 | fun apply f = applys (Seq.single o f); | 
| 6902 | 78 | |
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 79 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 80 | end; |