| author | wenzelm | 
| Sat, 01 Jun 2019 11:29:59 +0200 | |
| changeset 70299 | 83774d669b51 | 
| parent 56853 | a265e41cc33b | 
| child 73019 | 05e2cab9af8b | 
| permissions | -rw-r--r-- | 
| 31775 | 1  | 
(* Title: HOL/Tools/Function/scnp_solve.ML  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
2  | 
Author: Armin Heller, TU Muenchen  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
3  | 
Author: Alexander Krauss, TU Muenchen  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
4  | 
|
| 41114 | 5  | 
Certificate generation for SCNP using a SAT solver.  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
9  | 
signature SCNP_SOLVE =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
12  | 
datatype edge = GTR | GEQ  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
13  | 
datatype graph = G of int * int * (int * edge * int) list  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
14  | 
datatype graph_problem = GP of int list * graph list  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
15  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
16  | 
datatype label = MIN | MAX | MS  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
18  | 
type certificate =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
19  | 
label (* which order *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
20  | 
* (int * int) list list (* (multi)sets *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
21  | 
* int list (* strictly ordered calls *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
22  | 
* (int -> bool -> int -> (int * int) option) (* covering function *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
24  | 
val generate_certificate : bool -> label list -> graph_problem -> certificate option  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
25  | 
|
| 32740 | 26  | 
val solver : string Unsynchronized.ref  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
27  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
28  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
29  | 
structure ScnpSolve : SCNP_SOLVE =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
30  | 
struct  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
31  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
32  | 
(** Graph problems **)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
33  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
34  | 
datatype edge = GTR | GEQ ;  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
35  | 
datatype graph = G of int * int * (int * edge * int) list ;  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
36  | 
datatype graph_problem = GP of int list * graph list ;  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
37  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
38  | 
datatype label = MIN | MAX | MS ;  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
39  | 
type certificate =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
40  | 
label  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
41  | 
* (int * int) list list  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
42  | 
* int list  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
43  | 
* (int -> bool -> int -> (int * int) option)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
45  | 
fun graph_at (GP (_, gs), i) = nth gs i ;  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
46  | 
fun num_prog_pts (GP (arities, _)) = length arities ;  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
47  | 
fun num_graphs (GP (_, gs)) = length gs ;  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
48  | 
fun arity (GP (arities, gl)) i = nth arities i ;  | 
| 33002 | 49  | 
fun ndigits (GP (arities, _)) = IntInf.log2 (Integer.sum arities) + 1  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
52  | 
(** Propositional formulas **)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
53  | 
|
| 41471 | 54  | 
val Not = Prop_Logic.Not and And = Prop_Logic.And and Or = Prop_Logic.Or  | 
55  | 
val BoolVar = Prop_Logic.BoolVar  | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
56  | 
fun Implies (p, q) = Or (Not p, q)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
57  | 
fun Equiv (p, q) = And (Implies (p, q), Implies (q, p))  | 
| 41471 | 58  | 
val all = Prop_Logic.all  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
59  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
60  | 
(* finite indexed quantifiers:  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
61  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
62  | 
iforall n f <==> /\  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
63  | 
/ \ f i  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
64  | 
0<=i<n  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
65  | 
*)  | 
| 33063 | 66  | 
fun iforall n f = all (map_range f n)  | 
| 41471 | 67  | 
fun iexists n f = Prop_Logic.exists (map_range f n)  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
68  | 
fun iforall2 n m f = all (map_product f (0 upto n - 1) (0 upto m - 1))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
69  | 
|
| 33040 | 70  | 
fun the_one var n x = all (var x :: map (Not o var) (remove (op =) x (0 upto n - 1)))  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
71  | 
fun exactly_one n f = iexists n (the_one f n)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
73  | 
(* SAT solving *)  | 
| 32740 | 74  | 
val solver = Unsynchronized.ref "auto";  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
75  | 
fun sat_solver x =  | 
| 56853 | 76  | 
Function_Common.PROFILE "sat_solving..." (SAT_Solver.invoke_solver (!solver)) x  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
78  | 
(* "Virtual constructors" for various propositional variables *)  | 
| 
33855
 
cd8acf137c9c
eliminated dead code and some unused bindings, reported by polyml
 
krauss 
parents: 
33099 
diff
changeset
 | 
79  | 
fun var_constrs (gp as GP (arities, _)) =  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
80  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
81  | 
val n = Int.max (num_graphs gp, num_prog_pts gp)  | 
| 33029 | 82  | 
val k = fold Integer.max arities 1  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
83  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
84  | 
(* Injective, provided a < 8, x < n, and i < k. *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
85  | 
fun prod a x i j = ((j * k + i) * n + x) * 8 + a + 1  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
86  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
87  | 
fun ES (g, i, j) = BoolVar (prod 0 g i j)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
88  | 
fun EW (g, i, j) = BoolVar (prod 1 g i j)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
89  | 
fun WEAK g = BoolVar (prod 2 g 0 0)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
90  | 
fun STRICT g = BoolVar (prod 3 g 0 0)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
91  | 
fun P (p, i) = BoolVar (prod 4 p i 0)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
92  | 
fun GAM (g, i, j)= BoolVar (prod 5 g i j)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
93  | 
fun EPS (g, i) = BoolVar (prod 6 g i 0)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
94  | 
fun TAG (p, i) b = BoolVar (prod 7 p i b)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
95  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
96  | 
(ES,EW,WEAK,STRICT,P,GAM,EPS,TAG)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
97  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
98  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
100  | 
fun graph_info gp g =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
101  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
102  | 
val G (p, q, edgs) = graph_at (gp, g)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
103  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
104  | 
(g, p, q, arity gp p, arity gp q, edgs)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
105  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
107  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
108  | 
(* Order-independent part of encoding *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
110  | 
fun encode_graphs bits gp =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
111  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
112  | 
val ng = num_graphs gp  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
113  | 
val (ES,EW,_,_,_,_,_,TAG) = var_constrs gp  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
114  | 
|
| 41471 | 115  | 
fun encode_constraint_strict 0 (x, y) = Prop_Logic.False  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
116  | 
| encode_constraint_strict k (x, y) =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
117  | 
Or (And (TAG x (k - 1), Not (TAG y (k - 1))),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
118  | 
And (Equiv (TAG x (k - 1), TAG y (k - 1)),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
119  | 
encode_constraint_strict (k - 1) (x, y)))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
120  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
121  | 
fun encode_constraint_weak k (x, y) =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
122  | 
Or (encode_constraint_strict k (x, y),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
123  | 
iforall k (fn i => Equiv (TAG x i, TAG y i)))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
124  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
125  | 
fun encode_graph (g, p, q, n, m, edges) =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
126  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
127  | 
fun encode_edge i j =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
128  | 
if exists (fn x => x = (i, GTR, j)) edges then  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
129  | 
And (ES (g, i, j), EW (g, i, j))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
130  | 
else if not (exists (fn x => x = (i, GEQ, j)) edges) then  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
131  | 
And (Not (ES (g, i, j)), Not (EW (g, i, j)))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
132  | 
else  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
133  | 
And (  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
134  | 
Equiv (ES (g, i, j),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
135  | 
encode_constraint_strict bits ((p, i), (q, j))),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
136  | 
Equiv (EW (g, i, j),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
137  | 
encode_constraint_weak bits ((p, i), (q, j))))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
138  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
139  | 
iforall2 n m encode_edge  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
140  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
141  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
142  | 
iforall ng (encode_graph o graph_info gp)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
143  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
144  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
145  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
146  | 
(* Order-specific part of encoding *)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
147  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
148  | 
fun encode bits gp mu =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
149  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
150  | 
val ng = num_graphs gp  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
151  | 
val (ES,EW,WEAK,STRICT,P,GAM,EPS,_) = var_constrs gp  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
152  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
153  | 
fun encode_graph MAX (g, p, q, n, m, _) =  | 
| 29712 | 154  | 
And (  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
155  | 
Equiv (WEAK g,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
156  | 
iforall m (fn j =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
157  | 
Implies (P (q, j),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
158  | 
iexists n (fn i =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
159  | 
And (P (p, i), EW (g, i, j)))))),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
160  | 
Equiv (STRICT g,  | 
| 29712 | 161  | 
And (  | 
162  | 
iforall m (fn j =>  | 
|
163  | 
Implies (P (q, j),  | 
|
164  | 
iexists n (fn i =>  | 
|
165  | 
And (P (p, i), ES (g, i, j))))),  | 
|
166  | 
iexists n (fn i => P (p, i)))))  | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
167  | 
| encode_graph MIN (g, p, q, n, m, _) =  | 
| 29712 | 168  | 
And (  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
169  | 
Equiv (WEAK g,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
170  | 
iforall n (fn i =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
171  | 
Implies (P (p, i),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
172  | 
iexists m (fn j =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
173  | 
And (P (q, j), EW (g, i, j)))))),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
174  | 
Equiv (STRICT g,  | 
| 29712 | 175  | 
And (  | 
176  | 
iforall n (fn i =>  | 
|
177  | 
Implies (P (p, i),  | 
|
178  | 
iexists m (fn j =>  | 
|
179  | 
And (P (q, j), ES (g, i, j))))),  | 
|
180  | 
iexists m (fn j => P (q, j)))))  | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
181  | 
| encode_graph MS (g, p, q, n, m, _) =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
182  | 
all [  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
183  | 
Equiv (WEAK g,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
184  | 
iforall m (fn j =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
185  | 
Implies (P (q, j),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
186  | 
iexists n (fn i => GAM (g, i, j))))),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
187  | 
Equiv (STRICT g,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
188  | 
iexists n (fn i =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
189  | 
And (P (p, i), Not (EPS (g, i))))),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
190  | 
iforall2 n m (fn i => fn j =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
191  | 
Implies (GAM (g, i, j),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
192  | 
all [  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
193  | 
P (p, i),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
194  | 
P (q, j),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
195  | 
EW (g, i, j),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
196  | 
Equiv (Not (EPS (g, i)), ES (g, i, j))])),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
197  | 
iforall n (fn i =>  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
198  | 
Implies (And (P (p, i), EPS (g, i)),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
199  | 
exactly_one m (fn j => GAM (g, i, j))))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
200  | 
]  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
201  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
202  | 
all [  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
203  | 
encode_graphs bits gp,  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
204  | 
iforall ng (encode_graph mu o graph_info gp),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
205  | 
iforall ng (fn x => WEAK x),  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
206  | 
iexists ng (fn x => STRICT x)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
207  | 
]  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
208  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
209  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
210  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
211  | 
(*Generieren des level-mapping und diverser output*)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
212  | 
fun mk_certificate bits label gp f =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
213  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
214  | 
val (ES,EW,WEAK,STRICT,P,GAM,EPS,TAG) = var_constrs gp  | 
| 41471 | 215  | 
fun assign (Prop_Logic.BoolVar v) = the_default false (f v)  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
216  | 
fun assignTag i j =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
217  | 
(fold (fn x => fn y => 2 * y + (if assign (TAG (i, j) x) then 1 else 0))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
218  | 
(bits - 1 downto 0) 0)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
219  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
220  | 
val level_mapping =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
221  | 
let fun prog_pt_mapping p =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
222  | 
map_filter (fn x => if assign (P(p, x)) then SOME (x, assignTag p x) else NONE)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
223  | 
(0 upto (arity gp p) - 1)  | 
| 33063 | 224  | 
in map_range prog_pt_mapping (num_prog_pts gp) end  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
225  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
226  | 
val strict_list = filter (assign o STRICT) (0 upto num_graphs gp - 1)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
227  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
228  | 
fun covering_pair g bStrict j =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
229  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
230  | 
val (_, p, q, n, m, _) = graph_info gp g  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
231  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
232  | 
fun cover MAX j = find_index (fn i => assign (P (p, i)) andalso assign (EW (g, i, j))) (0 upto n - 1)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
233  | 
| cover MS k = find_index (fn i => assign (GAM (g, i, k))) (0 upto n - 1)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
234  | 
| cover MIN i = find_index (fn j => assign (P (q, j)) andalso assign (EW (g, i, j))) (0 upto m - 1)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
235  | 
fun cover_strict MAX j = find_index (fn i => assign (P (p, i)) andalso assign (ES (g, i, j))) (0 upto n - 1)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
236  | 
| cover_strict MS k = find_index (fn i => assign (GAM (g, i, k)) andalso not (assign (EPS (g, i) ))) (0 upto n - 1)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
237  | 
| cover_strict MIN i = find_index (fn j => assign (P (q, j)) andalso assign (ES (g, i, j))) (0 upto m - 1)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
238  | 
val i = if bStrict then cover_strict label j else cover label j  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
239  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
240  | 
find_first (fn x => fst x = i) (nth level_mapping (if label = MIN then q else p))  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
241  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
242  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
243  | 
(label, level_mapping, strict_list, covering_pair)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
244  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
245  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
246  | 
(*interface for the proof reconstruction*)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
247  | 
fun generate_certificate use_tags labels gp =  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
248  | 
let  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
249  | 
val bits = if use_tags then ndigits gp else 0  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
250  | 
in  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
251  | 
get_first  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
252  | 
(fn l => case sat_solver (encode bits gp l) of  | 
| 56853 | 253  | 
SAT_Solver.SATISFIABLE f => SOME (mk_certificate bits l gp f)  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
254  | 
| _ => NONE)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
255  | 
labels  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
256  | 
end  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents:  
diff
changeset
 | 
257  | 
end  |