| author | wenzelm | 
| Sat, 12 Nov 2011 20:14:09 +0100 | |
| changeset 45476 | 6f9e24376ffd | 
| parent 45398 | 7dbb7b044a11 | 
| child 55889 | 6bfbec3dff62 | 
| permissions | -rw-r--r-- | 
| 33982 | 1  | 
(* Title: HOL/Tools/Nitpick/nitpick_peephole.ML  | 
| 33192 | 2  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
34982
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34936 
diff
changeset
 | 
3  | 
Copyright 2008, 2009, 2010  | 
| 33192 | 4  | 
|
5  | 
Peephole optimizer for Nitpick.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature NITPICK_PEEPHOLE =  | 
|
9  | 
sig  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
10  | 
type n_ary_index = Kodkod.n_ary_index  | 
| 33192 | 11  | 
type formula = Kodkod.formula  | 
12  | 
type int_expr = Kodkod.int_expr  | 
|
13  | 
type rel_expr = Kodkod.rel_expr  | 
|
14  | 
type decl = Kodkod.decl  | 
|
15  | 
type expr_assign = Kodkod.expr_assign  | 
|
16  | 
||
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
17  | 
type name_pool =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
18  | 
    {rels: n_ary_index list,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
19  | 
vars: n_ary_index list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
20  | 
formula_reg: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
21  | 
rel_reg: int}  | 
| 33192 | 22  | 
|
23  | 
val initial_pool : name_pool  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
24  | 
val not3_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
25  | 
val suc_rel : n_ary_index  | 
| 38126 | 26  | 
val suc_rels_base : int  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
27  | 
val unsigned_bit_word_sel_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
28  | 
val signed_bit_word_sel_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
29  | 
val nat_add_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
30  | 
val int_add_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
31  | 
val nat_subtract_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
32  | 
val int_subtract_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
33  | 
val nat_multiply_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
34  | 
val int_multiply_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
35  | 
val nat_divide_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
36  | 
val int_divide_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
37  | 
val nat_less_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
38  | 
val int_less_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
39  | 
val gcd_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
40  | 
val lcm_rel : n_ary_index  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
41  | 
val norm_frac_rel : n_ary_index  | 
| 33192 | 42  | 
val atom_for_bool : int -> bool -> rel_expr  | 
43  | 
val formula_for_bool : bool -> formula  | 
|
44  | 
val atom_for_nat : int * int -> int -> int  | 
|
45  | 
val min_int_for_card : int -> int  | 
|
46  | 
val max_int_for_card : int -> int  | 
|
47  | 
val int_for_atom : int * int -> int -> int  | 
|
48  | 
val atom_for_int : int * int -> int -> int  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
49  | 
val is_twos_complement_representable : int -> int -> bool  | 
| 38126 | 50  | 
val suc_rel_for_atom_seq : (int * int) * bool -> n_ary_index  | 
51  | 
val atom_seq_for_suc_rel : n_ary_index -> (int * int) * bool  | 
|
| 33192 | 52  | 
val inline_rel_expr : rel_expr -> bool  | 
53  | 
val empty_n_ary_rel : int -> rel_expr  | 
|
54  | 
val num_seq : int -> int -> int_expr list  | 
|
55  | 
val s_and : formula -> formula -> formula  | 
|
56  | 
||
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
57  | 
type kodkod_constrs =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
58  | 
    {kk_all: decl list -> formula -> formula,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
59  | 
kk_exist: decl list -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
60  | 
kk_formula_let: expr_assign list -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
61  | 
kk_formula_if: formula -> formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
62  | 
kk_or: formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
63  | 
kk_not: formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
64  | 
kk_iff: formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
65  | 
kk_implies: formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
66  | 
kk_and: formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
67  | 
kk_subset: rel_expr -> rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
68  | 
kk_rel_eq: rel_expr -> rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
69  | 
kk_no: rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
70  | 
kk_lone: rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
71  | 
kk_one: rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
72  | 
kk_some: rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
73  | 
kk_rel_let: expr_assign list -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
74  | 
kk_rel_if: formula -> rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
75  | 
kk_union: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
76  | 
kk_difference: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
77  | 
kk_override: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
78  | 
kk_intersect: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
79  | 
kk_product: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
80  | 
kk_join: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
81  | 
kk_closure: rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
82  | 
kk_reflexive_closure: rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
83  | 
kk_comprehension: decl list -> formula -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
84  | 
kk_project: rel_expr -> int_expr list -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
85  | 
kk_project_seq: rel_expr -> int -> int -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
86  | 
kk_not3: rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
87  | 
kk_nat_less: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
88  | 
kk_int_less: rel_expr -> rel_expr -> rel_expr}  | 
| 33192 | 89  | 
|
90  | 
val kodkod_constrs : bool -> int -> int -> int -> kodkod_constrs  | 
|
91  | 
end;  | 
|
92  | 
||
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
93  | 
structure Nitpick_Peephole : NITPICK_PEEPHOLE =  | 
| 33192 | 94  | 
struct  | 
95  | 
||
96  | 
open Kodkod  | 
|
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
97  | 
open Nitpick_Util  | 
| 33192 | 98  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
99  | 
type name_pool =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
100  | 
  {rels: n_ary_index list,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
101  | 
vars: n_ary_index list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
102  | 
formula_reg: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
103  | 
rel_reg: int}  | 
| 33192 | 104  | 
|
| 38126 | 105  | 
(* FIXME: needed? *)  | 
106  | 
val initial_pool = {rels = [], vars = [], formula_reg = 10, rel_reg = 10}
 | 
|
| 33192 | 107  | 
|
| 38126 | 108  | 
val not3_rel = (2, ~1)  | 
109  | 
val unsigned_bit_word_sel_rel = (2, ~2)  | 
|
110  | 
val signed_bit_word_sel_rel = (2, ~3)  | 
|
111  | 
val suc_rel = (2, ~4)  | 
|
112  | 
val suc_rels_base = ~5 (* must be the last of the binary series *)  | 
|
113  | 
val nat_add_rel = (3, ~1)  | 
|
114  | 
val int_add_rel = (3, ~2)  | 
|
115  | 
val nat_subtract_rel = (3, ~3)  | 
|
116  | 
val int_subtract_rel = (3, ~4)  | 
|
117  | 
val nat_multiply_rel = (3, ~5)  | 
|
118  | 
val int_multiply_rel = (3, ~6)  | 
|
119  | 
val nat_divide_rel = (3, ~7)  | 
|
120  | 
val int_divide_rel = (3, ~8)  | 
|
121  | 
val nat_less_rel = (3, ~9)  | 
|
122  | 
val int_less_rel = (3, ~10)  | 
|
123  | 
val gcd_rel = (3, ~11)  | 
|
124  | 
val lcm_rel = (3, ~12)  | 
|
125  | 
val norm_frac_rel = (4, ~1)  | 
|
| 33192 | 126  | 
|
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35284 
diff
changeset
 | 
127  | 
fun atom_for_bool j0 = Atom o Integer.add j0 o int_from_bool  | 
| 33192 | 128  | 
fun formula_for_bool b = if b then True else False  | 
129  | 
||
130  | 
fun atom_for_nat (k, j0) n = if n < 0 orelse n >= k then ~1 else n + j0  | 
|
131  | 
fun min_int_for_card k = ~k div 2 + 1  | 
|
132  | 
fun max_int_for_card k = k div 2  | 
|
133  | 
fun int_for_atom (k, j0) j =  | 
|
134  | 
let val j = j - j0 in if j <= max_int_for_card k then j else j - k end  | 
|
135  | 
fun atom_for_int (k, j0) n =  | 
|
136  | 
if n < min_int_for_card k orelse n > max_int_for_card k then ~1  | 
|
137  | 
else if n < 0 then n + k + j0  | 
|
138  | 
else n + j0  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
139  | 
fun is_twos_complement_representable bits n =  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
140  | 
let val max = reasonable_power 2 bits in n >= ~ max andalso n < max end  | 
| 33192 | 141  | 
|
| 38126 | 142  | 
val max_squeeze_card = 49  | 
143  | 
||
144  | 
fun squeeze (m, n) =  | 
|
145  | 
if n > max_squeeze_card then  | 
|
146  | 
    raise TOO_LARGE ("Nitpick_Peephole.squeeze",
 | 
|
147  | 
                     "too large cardinality (" ^ string_of_int n ^ ")")
 | 
|
148  | 
else  | 
|
149  | 
(max_squeeze_card + 1) * m + n  | 
|
150  | 
fun unsqueeze p = (p div (max_squeeze_card + 1), p mod (max_squeeze_card + 1))  | 
|
151  | 
||
152  | 
fun boolify (j, b) = 2 * j + (if b then 0 else 1)  | 
|
153  | 
fun unboolify j = (j div 2, j mod 2 = 0)  | 
|
154  | 
||
155  | 
fun suc_rel_for_atom_seq (x, tabulate) =  | 
|
156  | 
(2, suc_rels_base - boolify (squeeze x, tabulate))  | 
|
157  | 
fun atom_seq_for_suc_rel (_, j) = unboolify (~ j + suc_rels_base) |>> unsqueeze  | 
|
158  | 
||
| 33192 | 159  | 
fun is_none_product (Product (r1, r2)) =  | 
160  | 
is_none_product r1 orelse is_none_product r2  | 
|
161  | 
| is_none_product None = true  | 
|
162  | 
| is_none_product _ = false  | 
|
163  | 
||
164  | 
fun is_one_rel_expr (Atom _) = true  | 
|
165  | 
| is_one_rel_expr (AtomSeq (1, _)) = true  | 
|
166  | 
| is_one_rel_expr (Var _) = true  | 
|
167  | 
| is_one_rel_expr _ = false  | 
|
168  | 
||
169  | 
fun inline_rel_expr (Product (r1, r2)) =  | 
|
170  | 
inline_rel_expr r1 andalso inline_rel_expr r2  | 
|
171  | 
| inline_rel_expr Iden = true  | 
|
172  | 
| inline_rel_expr Ints = true  | 
|
173  | 
| inline_rel_expr None = true  | 
|
174  | 
| inline_rel_expr Univ = true  | 
|
175  | 
| inline_rel_expr (Atom _) = true  | 
|
176  | 
| inline_rel_expr (AtomSeq _) = true  | 
|
177  | 
| inline_rel_expr (Rel _) = true  | 
|
178  | 
| inline_rel_expr (Var _) = true  | 
|
179  | 
| inline_rel_expr (RelReg _) = true  | 
|
180  | 
| inline_rel_expr _ = false  | 
|
181  | 
||
182  | 
fun rel_expr_equal None (Atom _) = SOME false  | 
|
183  | 
| rel_expr_equal None (AtomSeq (k, _)) = SOME (k = 0)  | 
|
184  | 
| rel_expr_equal (Atom _) None = SOME false  | 
|
185  | 
| rel_expr_equal (AtomSeq (k, _)) None = SOME (k = 0)  | 
|
186  | 
| rel_expr_equal (Atom j1) (Atom j2) = SOME (j1 = j2)  | 
|
187  | 
| rel_expr_equal (Atom j) (AtomSeq (k, j0)) = SOME (j = j0 andalso k = 1)  | 
|
188  | 
| rel_expr_equal (AtomSeq (k, j0)) (Atom j) = SOME (j = j0 andalso k = 1)  | 
|
189  | 
| rel_expr_equal (AtomSeq x1) (AtomSeq x2) = SOME (x1 = x2)  | 
|
190  | 
| rel_expr_equal r1 r2 = if r1 = r2 then SOME true else NONE  | 
|
191  | 
||
192  | 
fun rel_expr_intersects (Atom j1) (Atom j2) = SOME (j1 = j2)  | 
|
193  | 
| rel_expr_intersects (Atom j) (AtomSeq (k, j0)) = SOME (j < j0 + k)  | 
|
194  | 
| rel_expr_intersects (AtomSeq (k, j0)) (Atom j) = SOME (j < j0 + k)  | 
|
195  | 
| rel_expr_intersects (AtomSeq (k1, j01)) (AtomSeq (k2, j02)) =  | 
|
196  | 
SOME (k1 > 0 andalso k2 > 0 andalso j01 + k1 > j02 andalso j02 + k2 > j01)  | 
|
197  | 
| rel_expr_intersects r1 r2 =  | 
|
198  | 
if is_none_product r1 orelse is_none_product r2 then SOME false else NONE  | 
|
199  | 
||
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
200  | 
fun empty_n_ary_rel 0 = raise ARG ("Nitpick_Peephole.empty_n_ary_rel", "0")
 | 
| 33192 | 201  | 
| empty_n_ary_rel n = funpow (n - 1) (curry Product None) None  | 
202  | 
||
203  | 
fun decl_one_set (DeclOne (_, r)) = r  | 
|
204  | 
| decl_one_set _ =  | 
|
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
205  | 
    raise ARG ("Nitpick_Peephole.decl_one_set", "not \"DeclOne\"")
 | 
| 33192 | 206  | 
|
207  | 
fun is_Num (Num _) = true  | 
|
208  | 
| is_Num _ = false  | 
|
209  | 
fun dest_Num (Num k) = k  | 
|
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
210  | 
  | dest_Num _ = raise ARG ("Nitpick_Peephole.dest_Num", "not \"Num\"")
 | 
| 33192 | 211  | 
fun num_seq j0 n = map Num (index_seq j0 n)  | 
212  | 
||
213  | 
fun occurs_in_union r (Union (r1, r2)) =  | 
|
214  | 
occurs_in_union r r1 orelse occurs_in_union r r2  | 
|
215  | 
| occurs_in_union r r' = (r = r')  | 
|
216  | 
||
217  | 
fun s_and True f2 = f2  | 
|
218  | 
| s_and False _ = False  | 
|
219  | 
| s_and f1 True = f1  | 
|
220  | 
| s_and _ False = False  | 
|
221  | 
| s_and f1 f2 = And (f1, f2)  | 
|
222  | 
||
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
223  | 
type kodkod_constrs =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
224  | 
  {kk_all: decl list -> formula -> formula,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
225  | 
kk_exist: decl list -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
226  | 
kk_formula_let: expr_assign list -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
227  | 
kk_formula_if: formula -> formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
228  | 
kk_or: formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
229  | 
kk_not: formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
230  | 
kk_iff: formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
231  | 
kk_implies: formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
232  | 
kk_and: formula -> formula -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
233  | 
kk_subset: rel_expr -> rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
234  | 
kk_rel_eq: rel_expr -> rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
235  | 
kk_no: rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
236  | 
kk_lone: rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
237  | 
kk_one: rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
238  | 
kk_some: rel_expr -> formula,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
239  | 
kk_rel_let: expr_assign list -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
240  | 
kk_rel_if: formula -> rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
241  | 
kk_union: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
242  | 
kk_difference: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
243  | 
kk_override: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
244  | 
kk_intersect: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
245  | 
kk_product: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
246  | 
kk_join: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
247  | 
kk_closure: rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
248  | 
kk_reflexive_closure: rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
249  | 
kk_comprehension: decl list -> formula -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
250  | 
kk_project: rel_expr -> int_expr list -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
251  | 
kk_project_seq: rel_expr -> int -> int -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
252  | 
kk_not3: rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
253  | 
kk_nat_less: rel_expr -> rel_expr -> rel_expr,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36385 
diff
changeset
 | 
254  | 
kk_int_less: rel_expr -> rel_expr -> rel_expr}  | 
| 33192 | 255  | 
|
256  | 
(* We assume throughout that Kodkod variables have a "one" constraint. This is  | 
|
257  | 
always the case if Kodkod's skolemization is disabled. *)  | 
|
258  | 
fun kodkod_constrs optim nat_card int_card main_j0 =  | 
|
259  | 
let  | 
|
260  | 
val from_bool = atom_for_bool main_j0  | 
|
261  | 
fun from_nat n = Atom (n + main_j0)  | 
|
262  | 
fun to_nat j = j - main_j0  | 
|
263  | 
val to_int = int_for_atom (int_card, main_j0)  | 
|
264  | 
||
| 38126 | 265  | 
val exists_empty_decl = exists (fn DeclOne (_, None) => true | _ => false)  | 
266  | 
||
| 33192 | 267  | 
fun s_all _ True = True  | 
268  | 
| s_all _ False = False  | 
|
269  | 
| s_all [] f = f  | 
|
| 38126 | 270  | 
| s_all ds (All (ds', f)) = s_all (ds @ ds') f  | 
271  | 
| s_all ds f = if exists_empty_decl ds then True else All (ds, f)  | 
|
| 33192 | 272  | 
fun s_exist _ True = True  | 
273  | 
| s_exist _ False = False  | 
|
274  | 
| s_exist [] f = f  | 
|
| 38126 | 275  | 
| s_exist ds (Exist (ds', f)) = s_exist (ds @ ds') f  | 
276  | 
| s_exist ds f = if exists_empty_decl ds then False else Exist (ds, f)  | 
|
| 33192 | 277  | 
|
278  | 
fun s_formula_let _ True = True  | 
|
279  | 
| s_formula_let _ False = False  | 
|
280  | 
| s_formula_let assigns f = FormulaLet (assigns, f)  | 
|
281  | 
||
282  | 
fun s_not True = False  | 
|
283  | 
| s_not False = True  | 
|
284  | 
| s_not (All (ds, f)) = Exist (ds, s_not f)  | 
|
285  | 
| s_not (Exist (ds, f)) = All (ds, s_not f)  | 
|
286  | 
| s_not (Or (f1, f2)) = And (s_not f1, s_not f2)  | 
|
287  | 
| s_not (Implies (f1, f2)) = And (f1, s_not f2)  | 
|
288  | 
| s_not (And (f1, f2)) = Or (s_not f1, s_not f2)  | 
|
289  | 
| s_not (Not f) = f  | 
|
290  | 
| s_not (No r) = Some r  | 
|
291  | 
| s_not (Some r) = No r  | 
|
292  | 
| s_not f = Not f  | 
|
293  | 
||
294  | 
fun s_or True _ = True  | 
|
295  | 
| s_or False f2 = f2  | 
|
296  | 
| s_or _ True = True  | 
|
297  | 
| s_or f1 False = f1  | 
|
298  | 
| s_or f1 f2 = if f1 = f2 then f1 else Or (f1, f2)  | 
|
299  | 
fun s_iff True f2 = f2  | 
|
300  | 
| s_iff False f2 = s_not f2  | 
|
301  | 
| s_iff f1 True = f1  | 
|
302  | 
| s_iff f1 False = s_not f1  | 
|
303  | 
| s_iff f1 f2 = if f1 = f2 then True else Iff (f1, f2)  | 
|
304  | 
fun s_implies True f2 = f2  | 
|
305  | 
| s_implies False _ = True  | 
|
306  | 
| s_implies _ True = True  | 
|
307  | 
| s_implies f1 False = s_not f1  | 
|
308  | 
| s_implies f1 f2 = if f1 = f2 then True else Implies (f1, f2)  | 
|
309  | 
||
310  | 
fun s_formula_if True f2 _ = f2  | 
|
311  | 
| s_formula_if False _ f3 = f3  | 
|
312  | 
| s_formula_if f1 True f3 = s_or f1 f3  | 
|
313  | 
| s_formula_if f1 False f3 = s_and (s_not f1) f3  | 
|
314  | 
| s_formula_if f1 f2 True = s_implies f1 f2  | 
|
315  | 
| s_formula_if f1 f2 False = s_and f1 f2  | 
|
316  | 
| s_formula_if f f1 f2 = FormulaIf (f, f1, f2)  | 
|
317  | 
||
318  | 
fun s_project r is =  | 
|
319  | 
(case r of  | 
|
320  | 
Project (r1, is') =>  | 
|
321  | 
if forall is_Num is then  | 
|
322  | 
s_project r1 (map (nth is' o dest_Num) is)  | 
|
323  | 
else  | 
|
324  | 
raise SAME ()  | 
|
325  | 
| _ => raise SAME ())  | 
|
326  | 
handle SAME () =>  | 
|
327  | 
let val n = length is in  | 
|
328  | 
if arity_of_rel_expr r = n andalso is = num_seq 0 n then r  | 
|
329  | 
else Project (r, is)  | 
|
330  | 
end  | 
|
331  | 
||
| 
35284
 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 
blanchet 
parents: 
35280 
diff
changeset
 | 
332  | 
fun s_xone xone r =  | 
| 
 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 
blanchet 
parents: 
35280 
diff
changeset
 | 
333  | 
if is_one_rel_expr r then  | 
| 
 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 
blanchet 
parents: 
35280 
diff
changeset
 | 
334  | 
True  | 
| 
 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 
blanchet 
parents: 
35280 
diff
changeset
 | 
335  | 
else case arity_of_rel_expr r of  | 
| 
 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 
blanchet 
parents: 
35280 
diff
changeset
 | 
336  | 
1 => xone r  | 
| 
 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 
blanchet 
parents: 
35280 
diff
changeset
 | 
337  | 
| arity => foldl1 And (map (xone o s_project r o single o Num)  | 
| 
 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 
blanchet 
parents: 
35280 
diff
changeset
 | 
338  | 
(index_seq 0 arity))  | 
| 33192 | 339  | 
fun s_no None = True  | 
340  | 
| s_no (Product (r1, r2)) = s_or (s_no r1) (s_no r2)  | 
|
| 34126 | 341  | 
| s_no (Intersect (Closure (Rel x), Iden)) = Acyclic x  | 
| 33192 | 342  | 
| s_no r = if is_one_rel_expr r then False else No r  | 
343  | 
fun s_lone None = True  | 
|
| 
35284
 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 
blanchet 
parents: 
35280 
diff
changeset
 | 
344  | 
| s_lone r = s_xone Lone r  | 
| 33192 | 345  | 
fun s_one None = False  | 
| 
35284
 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 
blanchet 
parents: 
35280 
diff
changeset
 | 
346  | 
| s_one r = s_xone One r  | 
| 33192 | 347  | 
fun s_some None = False  | 
348  | 
| s_some (Atom _) = True  | 
|
349  | 
| s_some (Product (r1, r2)) = s_and (s_some r1) (s_some r2)  | 
|
350  | 
| s_some r = if is_one_rel_expr r then True else Some r  | 
|
351  | 
||
352  | 
fun s_not3 (Atom j) = Atom (if j = main_j0 then j + 1 else j - 1)  | 
|
353  | 
| s_not3 (r as Join (r1, r2)) =  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
354  | 
if r2 = Rel not3_rel then r1 else Join (r, Rel not3_rel)  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
355  | 
| s_not3 r = Join (r, Rel not3_rel)  | 
| 33192 | 356  | 
|
357  | 
fun s_rel_eq r1 r2 =  | 
|
358  | 
(case (r1, r2) of  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
359  | 
(Join (r11, Rel x), _) =>  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
360  | 
if x = not3_rel then s_rel_eq r11 (s_not3 r2) else raise SAME ()  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
361  | 
| (RelIf (f, r11, r12), _) =>  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
362  | 
if inline_rel_expr r2 then  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
363  | 
s_formula_if f (s_rel_eq r11 r2) (s_rel_eq r12 r2)  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
364  | 
else  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
365  | 
raise SAME ()  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
366  | 
| (_, RelIf (f, r21, r22)) =>  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
367  | 
if inline_rel_expr r1 then  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
368  | 
s_formula_if f (s_rel_eq r1 r21) (s_rel_eq r1 r22)  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
369  | 
else  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
370  | 
raise SAME ()  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
371  | 
| (RelLet (bs, r1'), Atom _) => s_formula_let bs (s_rel_eq r1' r2)  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
372  | 
| (Atom _, RelLet (bs, r2')) => s_formula_let bs (s_rel_eq r1 r2')  | 
| 33192 | 373  | 
| _ => raise SAME ())  | 
374  | 
handle SAME () =>  | 
|
375  | 
case rel_expr_equal r1 r2 of  | 
|
376  | 
SOME true => True  | 
|
377  | 
| SOME false => False  | 
|
378  | 
| NONE =>  | 
|
379  | 
case (r1, r2) of  | 
|
380  | 
(_, RelIf (f, r21, r22)) =>  | 
|
381  | 
if inline_rel_expr r1 then  | 
|
382  | 
s_formula_if f (s_rel_eq r1 r21) (s_rel_eq r1 r22)  | 
|
383  | 
else  | 
|
384  | 
RelEq (r1, r2)  | 
|
385  | 
| (RelIf (f, r11, r12), _) =>  | 
|
386  | 
if inline_rel_expr r2 then  | 
|
387  | 
s_formula_if f (s_rel_eq r11 r2) (s_rel_eq r12 r2)  | 
|
388  | 
else  | 
|
389  | 
RelEq (r1, r2)  | 
|
| 34126 | 390  | 
| (_, None) => s_no r1  | 
391  | 
| (None, _) => s_no r2  | 
|
| 33192 | 392  | 
| _ => RelEq (r1, r2)  | 
393  | 
fun s_subset (Atom j1) (Atom j2) = formula_for_bool (j1 = j2)  | 
|
394  | 
| s_subset (Atom j) (AtomSeq (k, j0)) =  | 
|
395  | 
formula_for_bool (j >= j0 andalso j < j0 + k)  | 
|
| 
35280
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
34982 
diff
changeset
 | 
396  | 
| s_subset (Union (r11, r12)) r2 =  | 
| 33192 | 397  | 
s_and (s_subset r11 r2) (s_subset r12 r2)  | 
398  | 
| s_subset r1 (r2 as Union (r21, r22)) =  | 
|
399  | 
if is_one_rel_expr r1 then  | 
|
400  | 
s_or (s_subset r1 r21) (s_subset r1 r22)  | 
|
401  | 
else  | 
|
| 
34936
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34126 
diff
changeset
 | 
402  | 
if s_subset r1 r21 = True orelse s_subset r1 r22 = True orelse  | 
| 
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34126 
diff
changeset
 | 
403  | 
r1 = r2 then  | 
| 33192 | 404  | 
True  | 
405  | 
else  | 
|
406  | 
Subset (r1, r2)  | 
|
407  | 
| s_subset r1 r2 =  | 
|
408  | 
if r1 = r2 orelse is_none_product r1 then True  | 
|
409  | 
else if is_none_product r2 then s_no r1  | 
|
410  | 
else if forall is_one_rel_expr [r1, r2] then s_rel_eq r1 r2  | 
|
411  | 
else Subset (r1, r2)  | 
|
412  | 
||
413  | 
fun s_rel_let [b as AssignRelReg (x', r')] (r as RelReg x) =  | 
|
414  | 
if x = x' then r' else RelLet ([b], r)  | 
|
415  | 
| s_rel_let bs r = RelLet (bs, r)  | 
|
416  | 
||
417  | 
fun s_rel_if f r1 r2 =  | 
|
418  | 
(case (f, r1, r2) of  | 
|
419  | 
(True, _, _) => r1  | 
|
420  | 
| (False, _, _) => r2  | 
|
421  | 
| (No r1', None, RelIf (One r2', r3', r4')) =>  | 
|
422  | 
if r1' = r2' andalso r2' = r3' then s_rel_if (Lone r1') r1' r4'  | 
|
423  | 
else raise SAME ()  | 
|
424  | 
| _ => raise SAME ())  | 
|
425  | 
handle SAME () => if r1 = r2 then r1 else RelIf (f, r1, r2)  | 
|
426  | 
||
427  | 
fun s_union r1 (Union (r21, r22)) = s_union (s_union r1 r21) r22  | 
|
428  | 
| s_union r1 r2 =  | 
|
429  | 
if is_none_product r1 then r2  | 
|
430  | 
else if is_none_product r2 then r1  | 
|
431  | 
else if r1 = r2 then r1  | 
|
432  | 
else if occurs_in_union r2 r1 then r1  | 
|
433  | 
else Union (r1, r2)  | 
|
434  | 
fun s_difference r1 r2 =  | 
|
435  | 
if is_none_product r1 orelse is_none_product r2 then r1  | 
|
436  | 
else if r1 = r2 then empty_n_ary_rel (arity_of_rel_expr r1)  | 
|
437  | 
else Difference (r1, r2)  | 
|
438  | 
fun s_override r1 r2 =  | 
|
439  | 
if is_none_product r2 then r1  | 
|
440  | 
else if is_none_product r1 then r2  | 
|
441  | 
else Override (r1, r2)  | 
|
442  | 
fun s_intersect r1 r2 =  | 
|
443  | 
case rel_expr_intersects r1 r2 of  | 
|
444  | 
SOME true => if r1 = r2 then r1 else Intersect (r1, r2)  | 
|
445  | 
| SOME false => empty_n_ary_rel (arity_of_rel_expr r1)  | 
|
446  | 
| NONE => if is_none_product r1 then r1  | 
|
447  | 
else if is_none_product r2 then r2  | 
|
448  | 
else Intersect (r1, r2)  | 
|
449  | 
fun s_product r1 r2 =  | 
|
450  | 
if is_none_product r1 then  | 
|
451  | 
Product (r1, empty_n_ary_rel (arity_of_rel_expr r2))  | 
|
452  | 
else if is_none_product r2 then  | 
|
453  | 
Product (empty_n_ary_rel (arity_of_rel_expr r1), r2)  | 
|
454  | 
else  | 
|
455  | 
Product (r1, r2)  | 
|
456  | 
fun s_join r1 (Product (Product (r211, r212), r22)) =  | 
|
457  | 
Product (s_join r1 (Product (r211, r212)), r22)  | 
|
458  | 
| s_join (Product (r11, Product (r121, r122))) r2 =  | 
|
459  | 
Product (r11, s_join (Product (r121, r122)) r2)  | 
|
460  | 
| s_join None r = empty_n_ary_rel (arity_of_rel_expr r - 1)  | 
|
461  | 
| s_join r None = empty_n_ary_rel (arity_of_rel_expr r - 1)  | 
|
462  | 
| s_join (Product (None, None)) r = empty_n_ary_rel (arity_of_rel_expr r)  | 
|
463  | 
| s_join r (Product (None, None)) = empty_n_ary_rel (arity_of_rel_expr r)  | 
|
464  | 
| s_join Iden r2 = r2  | 
|
465  | 
| s_join r1 Iden = r1  | 
|
466  | 
| s_join (Product (r1, r2)) Univ =  | 
|
467  | 
if arity_of_rel_expr r2 = 1 then r1  | 
|
468  | 
else Product (r1, s_join r2 Univ)  | 
|
469  | 
| s_join Univ (Product (r1, r2)) =  | 
|
470  | 
if arity_of_rel_expr r1 = 1 then r2  | 
|
471  | 
else Product (s_join Univ r1, r2)  | 
|
472  | 
| s_join r1 (r2 as Product (r21, r22)) =  | 
|
473  | 
if arity_of_rel_expr r1 = 1 then  | 
|
474  | 
case rel_expr_intersects r1 r21 of  | 
|
475  | 
SOME true => r22  | 
|
476  | 
| SOME false => empty_n_ary_rel (arity_of_rel_expr r2 - 1)  | 
|
477  | 
| NONE => Join (r1, r2)  | 
|
478  | 
else  | 
|
479  | 
Join (r1, r2)  | 
|
480  | 
| s_join (r1 as Product (r11, r12)) r2 =  | 
|
481  | 
if arity_of_rel_expr r2 = 1 then  | 
|
482  | 
case rel_expr_intersects r2 r12 of  | 
|
483  | 
SOME true => r11  | 
|
484  | 
| SOME false => empty_n_ary_rel (arity_of_rel_expr r1 - 1)  | 
|
485  | 
| NONE => Join (r1, r2)  | 
|
486  | 
else  | 
|
487  | 
Join (r1, r2)  | 
|
488  | 
| s_join r1 (r2 as RelIf (f, r21, r22)) =  | 
|
489  | 
if inline_rel_expr r1 then s_rel_if f (s_join r1 r21) (s_join r1 r22)  | 
|
490  | 
else Join (r1, r2)  | 
|
491  | 
| s_join (r1 as RelIf (f, r11, r12)) r2 =  | 
|
492  | 
if inline_rel_expr r2 then s_rel_if f (s_join r11 r2) (s_join r12 r2)  | 
|
493  | 
else Join (r1, r2)  | 
|
| 
35280
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
34982 
diff
changeset
 | 
494  | 
| s_join (r1 as Atom j1) (r2 as Rel (x as (2, _))) =  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
495  | 
if x = suc_rel then  | 
| 33192 | 496  | 
let val n = to_nat j1 + 1 in  | 
497  | 
if n < nat_card then from_nat n else None  | 
|
498  | 
end  | 
|
499  | 
else  | 
|
500  | 
Join (r1, r2)  | 
|
501  | 
| s_join r1 (r2 as Project (r21, Num k :: is)) =  | 
|
502  | 
if k = arity_of_rel_expr r21 - 1 andalso arity_of_rel_expr r1 = 1 then  | 
|
503  | 
s_project (s_join r21 r1) is  | 
|
504  | 
else  | 
|
505  | 
Join (r1, r2)  | 
|
| 
35280
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
34982 
diff
changeset
 | 
506  | 
| s_join r1 (Join (r21, r22 as Rel (x as (3, _)))) =  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
507  | 
((if x = nat_add_rel then  | 
| 33192 | 508  | 
case (r21, r1) of  | 
509  | 
(Atom j1, Atom j2) =>  | 
|
510  | 
let val n = to_nat j1 + to_nat j2 in  | 
|
511  | 
if n < nat_card then from_nat n else None  | 
|
512  | 
end  | 
|
513  | 
| (Atom j, r) =>  | 
|
514  | 
(case to_nat j of  | 
|
515  | 
0 => r  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
516  | 
| 1 => s_join r (Rel suc_rel)  | 
| 33192 | 517  | 
| _ => raise SAME ())  | 
518  | 
| (r, Atom j) =>  | 
|
519  | 
(case to_nat j of  | 
|
520  | 
0 => r  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
521  | 
| 1 => s_join r (Rel suc_rel)  | 
| 33192 | 522  | 
| _ => raise SAME ())  | 
523  | 
| _ => raise SAME ()  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
524  | 
else if x = nat_subtract_rel then  | 
| 33192 | 525  | 
case (r21, r1) of  | 
| 
33705
 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 
blanchet 
parents: 
33232 
diff
changeset
 | 
526  | 
(Atom j1, Atom j2) => from_nat (nat_minus (to_nat j1) (to_nat j2))  | 
| 33192 | 527  | 
| _ => raise SAME ()  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
528  | 
else if x = nat_multiply_rel then  | 
| 33192 | 529  | 
case (r21, r1) of  | 
530  | 
(Atom j1, Atom j2) =>  | 
|
531  | 
let val n = to_nat j1 * to_nat j2 in  | 
|
532  | 
if n < nat_card then from_nat n else None  | 
|
533  | 
end  | 
|
534  | 
| (Atom j, r) =>  | 
|
535  | 
(case to_nat j of 0 => Atom j | 1 => r | _ => raise SAME ())  | 
|
536  | 
| (r, Atom j) =>  | 
|
537  | 
(case to_nat j of 0 => Atom j | 1 => r | _ => raise SAME ())  | 
|
538  | 
| _ => raise SAME ()  | 
|
539  | 
else  | 
|
540  | 
raise SAME ())  | 
|
541  | 
handle SAME () => List.foldr Join r22 [r1, r21])  | 
|
542  | 
| s_join r1 r2 = Join (r1, r2)  | 
|
543  | 
||
544  | 
fun s_closure Iden = Iden  | 
|
545  | 
| s_closure r = if is_none_product r then r else Closure r  | 
|
546  | 
fun s_reflexive_closure Iden = Iden  | 
|
547  | 
| s_reflexive_closure r =  | 
|
548  | 
if is_none_product r then Iden else ReflexiveClosure r  | 
|
549  | 
||
550  | 
fun s_comprehension ds False = empty_n_ary_rel (length ds)  | 
|
551  | 
| s_comprehension ds True = fold1 s_product (map decl_one_set ds)  | 
|
552  | 
| s_comprehension [d as DeclOne ((1, j1), r)]  | 
|
553  | 
(f as RelEq (Var (1, j2), Atom j)) =  | 
|
554  | 
if j1 = j2 andalso rel_expr_intersects (Atom j) r = SOME true then  | 
|
555  | 
Atom j  | 
|
556  | 
else  | 
|
557  | 
Comprehension ([d], f)  | 
|
558  | 
| s_comprehension ds f = Comprehension (ds, f)  | 
|
559  | 
||
560  | 
fun s_project_seq r =  | 
|
561  | 
let  | 
|
562  | 
fun aux arity r j0 n =  | 
|
563  | 
if j0 = 0 andalso arity = n then  | 
|
564  | 
r  | 
|
565  | 
else case r of  | 
|
566  | 
RelIf (f, r1, r2) =>  | 
|
567  | 
s_rel_if f (aux arity r1 j0 n) (aux arity r2 j0 n)  | 
|
568  | 
| Product (r1, r2) =>  | 
|
569  | 
let  | 
|
570  | 
val arity2 = arity_of_rel_expr r2  | 
|
571  | 
val arity1 = arity - arity2  | 
|
| 
33705
 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 
blanchet 
parents: 
33232 
diff
changeset
 | 
572  | 
val n1 = Int.min (nat_minus arity1 j0, n)  | 
| 33192 | 573  | 
val n2 = n - n1  | 
574  | 
fun one () = aux arity1 r1 j0 n1  | 
|
| 
33705
 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 
blanchet 
parents: 
33232 
diff
changeset
 | 
575  | 
fun two () = aux arity2 r2 (nat_minus j0 arity1) n2  | 
| 33192 | 576  | 
in  | 
577  | 
case (n1, n2) of  | 
|
578  | 
(0, _) => s_rel_if (s_some r1) (two ()) (empty_n_ary_rel n2)  | 
|
579  | 
| (_, 0) => s_rel_if (s_some r2) (one ()) (empty_n_ary_rel n1)  | 
|
580  | 
| _ => s_product (one ()) (two ())  | 
|
581  | 
end  | 
|
582  | 
| _ => s_project r (num_seq j0 n)  | 
|
583  | 
in aux (arity_of_rel_expr r) r end  | 
|
584  | 
||
585  | 
fun s_nat_less (Atom j1) (Atom j2) = from_bool (j1 < j2)  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
586  | 
| s_nat_less r1 r2 = fold s_join [r1, r2] (Rel nat_less_rel)  | 
| 33192 | 587  | 
fun s_int_less (Atom j1) (Atom j2) = from_bool (to_int j1 < to_int j2)  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
588  | 
| s_int_less r1 r2 = fold s_join [r1, r2] (Rel int_less_rel)  | 
| 33192 | 589  | 
|
590  | 
fun d_project_seq r j0 n = Project (r, num_seq j0 n)  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
591  | 
fun d_not3 r = Join (r, Rel not3_rel)  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
592  | 
fun d_nat_less r1 r2 = List.foldl Join (Rel nat_less_rel) [r1, r2]  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
33982 
diff
changeset
 | 
593  | 
fun d_int_less r1 r2 = List.foldl Join (Rel int_less_rel) [r1, r2]  | 
| 33192 | 594  | 
in  | 
595  | 
if optim then  | 
|
596  | 
      {kk_all = s_all, kk_exist = s_exist, kk_formula_let = s_formula_let,
 | 
|
597  | 
kk_formula_if = s_formula_if, kk_or = s_or, kk_not = s_not,  | 
|
598  | 
kk_iff = s_iff, kk_implies = s_implies, kk_and = s_and,  | 
|
599  | 
kk_subset = s_subset, kk_rel_eq = s_rel_eq, kk_no = s_no,  | 
|
600  | 
kk_lone = s_lone, kk_one = s_one, kk_some = s_some,  | 
|
601  | 
kk_rel_let = s_rel_let, kk_rel_if = s_rel_if, kk_union = s_union,  | 
|
602  | 
kk_difference = s_difference, kk_override = s_override,  | 
|
603  | 
kk_intersect = s_intersect, kk_product = s_product, kk_join = s_join,  | 
|
604  | 
kk_closure = s_closure, kk_reflexive_closure = s_reflexive_closure,  | 
|
605  | 
kk_comprehension = s_comprehension, kk_project = s_project,  | 
|
606  | 
kk_project_seq = s_project_seq, kk_not3 = s_not3,  | 
|
607  | 
kk_nat_less = s_nat_less, kk_int_less = s_int_less}  | 
|
608  | 
else  | 
|
609  | 
      {kk_all = curry All, kk_exist = curry Exist,
 | 
|
610  | 
kk_formula_let = curry FormulaLet, kk_formula_if = curry3 FormulaIf,  | 
|
611  | 
kk_or = curry Or,kk_not = Not, kk_iff = curry Iff, kk_implies = curry  | 
|
612  | 
Implies, kk_and = curry And, kk_subset = curry Subset, kk_rel_eq = curry  | 
|
613  | 
RelEq, kk_no = No, kk_lone = Lone, kk_one = One, kk_some = Some,  | 
|
614  | 
kk_rel_let = curry RelLet, kk_rel_if = curry3 RelIf, kk_union = curry  | 
|
615  | 
Union, kk_difference = curry Difference, kk_override = curry Override,  | 
|
616  | 
kk_intersect = curry Intersect, kk_product = curry Product,  | 
|
617  | 
kk_join = curry Join, kk_closure = Closure,  | 
|
618  | 
kk_reflexive_closure = ReflexiveClosure, kk_comprehension = curry  | 
|
619  | 
Comprehension, kk_project = curry Project,  | 
|
620  | 
kk_project_seq = d_project_seq, kk_not3 = d_not3,  | 
|
621  | 
kk_nat_less = d_nat_less, kk_int_less = d_int_less}  | 
|
622  | 
end  | 
|
623  | 
||
624  | 
end;  |