src/Pure/Isar/proof_history.ML
author haftmann
Mon, 17 Dec 2007 22:40:14 +0100
changeset 25684 2b3cce7d22b8
parent 23362 de1476695aa6
permissions -rw-r--r--
note in target
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
wenzelm
parents: 15531
diff changeset
     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
5f126c495771 removed nesting (unused);
wenzelm
parents: 6684
diff changeset
    11
  val position: T -> int
6684
4f859545bd92 adapted to History changes;
wenzelm
parents: 6374
diff changeset
    12
  val init: int option -> Proof.state -> T
4f859545bd92 adapted to History changes;
wenzelm
parents: 6374
diff changeset
    13
  val is_initial: T -> bool
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    14
  val map_current: (Proof.state -> Proof.state) -> T -> T
5827
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    15
  val current: T -> Proof.state
16814
b829a6c9a87a export previous;
wenzelm
parents: 16490
diff changeset
    16
  val previous: T -> Proof.state option
7731
51d59734743d clear: int argument;
wenzelm
parents: 7366
diff changeset
    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
17073
dc1040419645 back: removed ill-defined '!' option;
wenzelm
parents: 16814
diff changeset
    20
  val back: 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 =
16490
wenzelm
parents: 15531
diff changeset
    32
  Proof.state *                 (*first result*)
wenzelm
parents: 15531
diff changeset
    33
  Proof.state Seq.seq;          (*alternative results*)
5827
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    34
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    35
type proof = node * int;        (*proof node, proof position*)
5827
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
fun app f (ProofHistory x) = ProofHistory (f x);
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    40
fun hist_app f = app (History.apply f);
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    41
fun hist_map f = app (History.map_current f);
5827
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    42
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    43
fun position (ProofHistory prf) = snd (History.current prf);
6902
5f126c495771 removed nesting (unused);
wenzelm
parents: 6684
diff changeset
    44
5827
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
(* generic history operations *)
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    47
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    48
fun init lim st = ProofHistory (History.init lim ((st, Seq.empty), 0));
6684
4f859545bd92 adapted to History changes;
wenzelm
parents: 6374
diff changeset
    49
fun is_initial (ProofHistory prf) = History.is_initial prf;
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    50
fun map_current f = hist_map (apfst (apfst f));
6902
5f126c495771 removed nesting (unused);
wenzelm
parents: 6684
diff changeset
    51
fun current (ProofHistory prf) = fst (fst (History.current prf));
16814
b829a6c9a87a export previous;
wenzelm
parents: 16490
diff changeset
    52
fun previous (ProofHistory prf) = Option.map (fst o fst) (History.previous prf);
7731
51d59734743d clear: int argument;
wenzelm
parents: 7366
diff changeset
    53
val clear = app o History.clear;
5827
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    54
val undo = app History.undo;
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    55
val redo = app History.redo;
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
17073
dc1040419645 back: removed ill-defined '!' option;
wenzelm
parents: 16814
diff changeset
    58
(* backtracking *)
5827
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    59
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    60
val back = hist_app (fn ((_, stq), position) =>
5827
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    61
  (case Seq.pull stq of
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 17073
diff changeset
    62
    NONE => error "back: no alternatives"
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    63
  | SOME result => (result, position)));
5827
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    64
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
(* apply transformer *)
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    67
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    68
fun applys f = hist_app (fn (node as (st, _), position) =>
5827
77071ac7c7b5 Histories of proof states, with undo / redo and prev / back.
wenzelm
parents:
diff changeset
    69
  (case Seq.pull (f st) of
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 17073
diff changeset
    70
    NONE => error "empty result sequence -- proof command failed"
23362
de1476695aa6 added map_current;
wenzelm
parents: 18678
diff changeset
    71
  | SOME results => (results, position + 1)));
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
fun apply f = applys (Seq.single o f);
6902
5f126c495771 removed nesting (unused);
wenzelm
parents: 6684
diff changeset
    74
5827
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
end;