# HG changeset patch # User wenzelm # Date 1181685725 -7200 # Node ID de1476695aa6f0809222771a63e22d3157f176d8 # Parent df3d21caad2c4cc1e24da45a451c1c258573cd9f added map_current; simplified internal type proof; diff -r df3d21caad2c -r de1476695aa6 src/Pure/Isar/proof_history.ML --- a/src/Pure/Isar/proof_history.ML Wed Jun 13 00:02:04 2007 +0200 +++ b/src/Pure/Isar/proof_history.ML Wed Jun 13 00:02:05 2007 +0200 @@ -11,6 +11,7 @@ val position: T -> int val init: int option -> Proof.state -> T val is_initial: T -> bool + val map_current: (Proof.state -> Proof.state) -> T -> T val current: T -> Proof.state val previous: T -> Proof.state option val clear: int -> T -> T @@ -31,20 +32,22 @@ Proof.state * (*first result*) Proof.state Seq.seq; (*alternative results*) -type proof = node * node list; +type proof = node * int; (*proof node, proof position*) datatype T = ProofHistory of proof History.T; fun app f (ProofHistory x) = ProofHistory (f x); fun hist_app f = app (History.apply f); +fun hist_map f = app (History.map_current f); -fun position (ProofHistory prf) = length (snd (History.current prf)); +fun position (ProofHistory prf) = snd (History.current prf); (* generic history operations *) -fun init lim st = ProofHistory (History.init lim ((st, Seq.empty), [])); +fun init lim st = ProofHistory (History.init lim ((st, Seq.empty), 0)); fun is_initial (ProofHistory prf) = History.is_initial prf; +fun map_current f = hist_map (apfst (apfst f)); fun current (ProofHistory prf) = fst (fst (History.current prf)); fun previous (ProofHistory prf) = Option.map (fst o fst) (History.previous prf); val clear = app o History.clear; @@ -54,18 +57,18 @@ (* backtracking *) -val back = hist_app (fn ((_, stq), nodes) => +val back = hist_app (fn ((_, stq), position) => (case Seq.pull stq of NONE => error "back: no alternatives" - | SOME result => (result, nodes))); + | SOME result => (result, position))); (* apply transformer *) -fun applys f = hist_app (fn (node as (st, _), nodes) => +fun applys f = hist_app (fn (node as (st, _), position) => (case Seq.pull (f st) of NONE => error "empty result sequence -- proof command failed" - | SOME results => (results, node :: nodes))); + | SOME results => (results, position + 1))); fun apply f = applys (Seq.single o f);