src/Pure/Isar/proof_history.ML
changeset 23362 de1476695aa6
parent 18678 dd0c569fa43d
equal deleted inserted replaced
23361:df3d21caad2c 23362:de1476695aa6
     9 sig
     9 sig
    10   type T
    10   type T
    11   val position: T -> int
    11   val position: T -> int
    12   val init: int option -> Proof.state -> T
    12   val init: int option -> Proof.state -> T
    13   val is_initial: T -> bool
    13   val is_initial: T -> bool
       
    14   val map_current: (Proof.state -> Proof.state) -> T -> T
    14   val current: T -> Proof.state
    15   val current: T -> Proof.state
    15   val previous: T -> Proof.state option
    16   val previous: T -> Proof.state option
    16   val clear: int -> T -> T
    17   val clear: int -> T -> T
    17   val undo: T -> T
    18   val undo: T -> T
    18   val redo: T -> T
    19   val redo: T -> T
    29 
    30 
    30 type node =
    31 type node =
    31   Proof.state *                 (*first result*)
    32   Proof.state *                 (*first result*)
    32   Proof.state Seq.seq;          (*alternative results*)
    33   Proof.state Seq.seq;          (*alternative results*)
    33 
    34 
    34 type proof = node * node list;
    35 type proof = node * int;        (*proof node, proof position*)
    35 
    36 
    36 datatype T = ProofHistory of proof History.T;
    37 datatype T = ProofHistory of proof History.T;
    37 
    38 
    38 fun app f (ProofHistory x) = ProofHistory (f x);
    39 fun app f (ProofHistory x) = ProofHistory (f x);
    39 fun hist_app f = app (History.apply f);
    40 fun hist_app f = app (History.apply f);
       
    41 fun hist_map f = app (History.map_current f);
    40 
    42 
    41 fun position (ProofHistory prf) = length (snd (History.current prf));
    43 fun position (ProofHistory prf) = snd (History.current prf);
    42 
    44 
    43 
    45 
    44 (* generic history operations *)
    46 (* generic history operations *)
    45 
    47 
    46 fun init lim st = ProofHistory (History.init lim ((st, Seq.empty), []));
    48 fun init lim st = ProofHistory (History.init lim ((st, Seq.empty), 0));
    47 fun is_initial (ProofHistory prf) = History.is_initial prf;
    49 fun is_initial (ProofHistory prf) = History.is_initial prf;
       
    50 fun map_current f = hist_map (apfst (apfst f));
    48 fun current (ProofHistory prf) = fst (fst (History.current prf));
    51 fun current (ProofHistory prf) = fst (fst (History.current prf));
    49 fun previous (ProofHistory prf) = Option.map (fst o fst) (History.previous prf);
    52 fun previous (ProofHistory prf) = Option.map (fst o fst) (History.previous prf);
    50 val clear = app o History.clear;
    53 val clear = app o History.clear;
    51 val undo = app History.undo;
    54 val undo = app History.undo;
    52 val redo = app History.redo;
    55 val redo = app History.redo;
    53 
    56 
    54 
    57 
    55 (* backtracking *)
    58 (* backtracking *)
    56 
    59 
    57 val back = hist_app (fn ((_, stq), nodes) =>
    60 val back = hist_app (fn ((_, stq), position) =>
    58   (case Seq.pull stq of
    61   (case Seq.pull stq of
    59     NONE => error "back: no alternatives"
    62     NONE => error "back: no alternatives"
    60   | SOME result => (result, nodes)));
    63   | SOME result => (result, position)));
    61 
    64 
    62 
    65 
    63 (* apply transformer *)
    66 (* apply transformer *)
    64 
    67 
    65 fun applys f = hist_app (fn (node as (st, _), nodes) =>
    68 fun applys f = hist_app (fn (node as (st, _), position) =>
    66   (case Seq.pull (f st) of
    69   (case Seq.pull (f st) of
    67     NONE => error "empty result sequence -- proof command failed"
    70     NONE => error "empty result sequence -- proof command failed"
    68   | SOME results => (results, node :: nodes)));
    71   | SOME results => (results, position + 1)));
    69 
    72 
    70 fun apply f = applys (Seq.single o f);
    73 fun apply f = applys (Seq.single o f);
    71 
    74 
    72 
    75 
    73 end;
    76 end;