| author | wenzelm | 
| Sun, 09 Apr 2006 18:51:22 +0200 | |
| changeset 19387 | 6af442fa80c3 | 
| parent 18678 | dd0c569fa43d | 
| child 23362 | de1476695aa6 | 
| 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 | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 4 | |
| 16490 | 5 | Histories of proof states, with undo / redo and backtracking. | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 6 | *) | 
| 
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 | signature PROOF_HISTORY = | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 9 | sig | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 10 | type T | 
| 6902 | 11 | val position: T -> int | 
| 6684 | 12 | val init: int option -> Proof.state -> T | 
| 13 | val is_initial: T -> bool | |
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 14 | val current: T -> Proof.state | 
| 16814 | 15 | val previous: T -> Proof.state option | 
| 7731 | 16 | val clear: int -> T -> T | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 17 | val undo: T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 18 | val redo: T -> T | 
| 17073 | 19 | val back: T -> T | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 20 | 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 | 21 | val apply: (Proof.state -> Proof.state) -> T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 22 | end; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 23 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 24 | structure ProofHistory: PROOF_HISTORY = | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 25 | struct | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 26 | |
| 
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 | (* datatype proof history *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 29 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 30 | type node = | 
| 16490 | 31 | Proof.state * (*first result*) | 
| 32 | Proof.state Seq.seq; (*alternative results*) | |
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 33 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 34 | type proof = node * node list; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 35 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 36 | datatype T = ProofHistory of proof History.T; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 37 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 38 | fun app f (ProofHistory x) = ProofHistory (f x); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 39 | fun hist_app f = app (History.apply f); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 40 | |
| 6902 | 41 | fun position (ProofHistory prf) = length (snd (History.current prf)); | 
| 42 | ||
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 43 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 44 | (* generic history operations *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 45 | |
| 6902 | 46 | fun init lim st = ProofHistory (History.init lim ((st, Seq.empty), [])); | 
| 6684 | 47 | fun is_initial (ProofHistory prf) = History.is_initial prf; | 
| 6902 | 48 | fun current (ProofHistory prf) = fst (fst (History.current prf)); | 
| 16814 | 49 | fun previous (ProofHistory prf) = Option.map (fst o fst) (History.previous prf); | 
| 7731 | 50 | val clear = app o History.clear; | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 51 | val undo = app History.undo; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 52 | val redo = app History.redo; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 53 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 54 | |
| 17073 | 55 | (* backtracking *) | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 56 | |
| 17073 | 57 | val back = hist_app (fn ((_, stq), nodes) => | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 58 | (case Seq.pull stq of | 
| 18678 | 59 | NONE => error "back: no alternatives" | 
| 17073 | 60 | | SOME result => (result, nodes))); | 
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 61 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 62 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 63 | (* apply transformer *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 64 | |
| 6902 | 65 | 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 | 66 | (case Seq.pull (f st) of | 
| 18678 | 67 | NONE => error "empty result sequence -- proof command failed" | 
| 15531 | 68 | | SOME results => (results, node :: nodes))); | 
| 5827 
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 | fun apply f = applys (Seq.single o f); | 
| 6902 | 71 | |
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 72 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 73 | end; |