| author | wenzelm | 
| Thu, 11 Feb 1999 15:30:10 +0100 | |
| changeset 6269 | dbb48b0744d3 | 
| parent 5944 | dcc446da8e19 | 
| child 6374 | a67e4729efb2 | 
| 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 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 5 | 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 | 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 | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 11 | exception FAIL of string | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 12 | val init: Proof.state -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 13 | val current: T -> Proof.state | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 14 | val clear: T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 15 | val undo: T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 16 | val redo: T -> T | 
| 5944 | 17 | val undos: int -> T -> T | 
| 18 | val redos: int -> T -> T | |
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 19 | val position: T -> int | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 20 | val nesting: T -> int | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 21 | val prev: T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 22 | val up: T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 23 | val top: T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 24 | val back: T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 25 | 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 | 26 | val applys_open: (Proof.state -> Proof.state Seq.seq) -> T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 27 | val applys_close: (Proof.state -> Proof.state Seq.seq) -> T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 28 | val apply: (Proof.state -> Proof.state) -> T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 29 | val apply_open: (Proof.state -> Proof.state) -> T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 30 | val apply_close: (Proof.state -> Proof.state) -> T -> T | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 31 | end; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 32 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 33 | structure ProofHistory: PROOF_HISTORY = | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 34 | struct | 
| 
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 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 37 | (* datatype proof history *) | 
| 
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 | type node = | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 40 | (Proof.state * (*first result*) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 41 | Proof.state Seq.seq) * (*alternative results*) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 42 | int; (*nest level*) | 
| 
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 | type proof = node * node list; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 45 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 46 | datatype T = ProofHistory of proof History.T; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 47 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 48 | exception FAIL of string; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 49 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 50 | fun app f (ProofHistory x) = ProofHistory (f x); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 51 | fun hist_app f = app (History.apply f); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 52 | |
| 
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 | (* generic history operations *) | 
| 
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 | fun init st = ProofHistory (History.init (((st, Seq.empty), 0), [])); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 57 | fun current (ProofHistory prf) = fst (fst (fst (History.current prf))); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 58 | val clear = app History.clear; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 59 | val undo = app History.undo; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 60 | val redo = app History.redo; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 61 | |
| 5944 | 62 | fun undos n = funpow n undo; | 
| 63 | fun redos n = funpow n redo; | |
| 64 | ||
| 5827 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 65 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 66 | (* navigate *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 67 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 68 | fun position (ProofHistory prf) = length (snd (History.current prf)); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 69 | fun nesting (ProofHistory prf) = snd (fst (History.current prf)); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 70 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 71 | fun prev_until msg _ (_, []) = raise FAIL msg | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 72 | | prev_until msg pred (_, node :: nodes) = | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 73 | if pred node then (node, nodes) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 74 | else prev_until msg pred (node, nodes); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 75 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 76 | val prev = hist_app (prev_until "no previous proof state" (K true)); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 77 | fun up prf = hist_app (prev_until "no upper proof state" (fn (_, m) => m < nesting prf)) prf; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 78 | val top = hist_app (fn (node, nodes) => (last_elem (node :: nodes), [])); | 
| 
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 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 81 | (* backtrack *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 82 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 83 | fun back_fun (((_, stq), n), nodes) = | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 84 | (case Seq.pull stq of | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 85 | None => | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 86 | (case nodes of | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 87 | [] => raise FAIL "back: no alternatives" | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 88 | | nd :: nds => (writeln "back: trying previous proof state ..."; back_fun (nd, nds))) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 89 | | Some result => ((result, n), nodes)); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 90 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 91 | val back = hist_app back_fun; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 92 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 93 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 94 | (* apply transformer *) | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 95 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 96 | fun gen_apply d f (node as ((st, _), n): node, nodes) = | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 97 | (case Seq.pull (f st) of | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 98 | None => raise FAIL "empty result sequence -- command failed" | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 99 | | Some results => ((results, n + d), node :: nodes)); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 100 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 101 | val applys = hist_app o gen_apply 0; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 102 | val applys_open = hist_app o gen_apply 1; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 103 | val applys_close = hist_app o gen_apply ~1; | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 104 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 105 | fun apply f = applys (Seq.single o f); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 106 | fun apply_open f = applys_open (Seq.single o f); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 107 | fun apply_close f = applys_close (Seq.single o f); | 
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 108 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 109 | |
| 
77071ac7c7b5
Histories of proof states, with undo / redo and prev / back.
 wenzelm parents: diff
changeset | 110 | end; |