| author | wenzelm | 
| Tue, 15 Jul 2008 15:59:49 +0200 | |
| changeset 27608 | 8fd5662ccd97 | 
| parent 27561 | a928e3439067 | 
| child 29606 | fedb8be05f24 | 
| permissions | -rw-r--r-- | 
| 27561 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 1 | (* Title: Pure/Isar/proof_node.ML | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 3 | Author: Makarius | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 4 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 5 | Proof nodes with linear position and backtracking. | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 6 | *) | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 7 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 8 | signature PROOF_NODE = | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 9 | sig | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 10 | type T | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 11 | val init: Proof.state -> T | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 12 | val current: T -> Proof.state | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 13 | val position: T -> int | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 14 | val back: T -> T | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 15 | val applys: (Proof.state -> Proof.state Seq.seq) -> T -> T | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 16 | val apply: (Proof.state -> Proof.state) -> T -> T | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 17 | end; | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 18 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 19 | structure ProofNode: PROOF_NODE = | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 20 | struct | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 21 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 22 | (* datatype *) | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 23 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 24 | datatype T = ProofNode of | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 25 | (Proof.state * (*first result*) | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 26 | Proof.state Seq.seq) * (*alternative results*) | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 27 | int; (*linear proof position*) | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 28 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 29 | fun init st = ProofNode ((st, Seq.empty), 0); | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 30 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 31 | fun current (ProofNode ((st, _), _)) = st; | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 32 | fun position (ProofNode (_, n)) = n; | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 33 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 34 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 35 | (* backtracking *) | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 36 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 37 | fun back (ProofNode ((_, stq), n)) = | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 38 | (case Seq.pull stq of | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 39 | NONE => error "back: no alternatives" | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 40 | | SOME res => ProofNode (res, n)); | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 41 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 42 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 43 | (* apply transformer *) | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 44 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 45 | fun applys f (ProofNode (node as (st, _), n)) = | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 46 | (case Seq.pull (f st) of | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 47 | NONE => error "empty result sequence -- proof command failed" | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 48 | | SOME res => ProofNode (res, n + 1)); | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 49 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 50 | fun apply f = applys (Seq.single o f); | 
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 51 | |
| 
a928e3439067
Proof nodes with linear position and backtracking (replaces obsolete proof_history.ML).
 wenzelm parents: diff
changeset | 52 | end; |