| author | boehmes | 
| Sun, 19 Dec 2010 17:55:56 +0100 | |
| changeset 41280 | a7de9d36f4f2 | 
| parent 38126 | 8031d099379a | 
| child 45398 | 7dbb7b044a11 | 
| 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: 
34936diff
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: 
33982diff
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: 
36385diff
changeset | 17 | type name_pool = | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 18 |     {rels: n_ary_index list,
 | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 19 | vars: n_ary_index list, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 20 | formula_reg: int, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
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: 
33982diff
changeset | 24 | val not3_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
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: 
33982diff
changeset | 27 | val unsigned_bit_word_sel_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 28 | val signed_bit_word_sel_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 29 | val nat_add_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 30 | val int_add_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 31 | val nat_subtract_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 32 | val int_subtract_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 33 | val nat_multiply_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 34 | val int_multiply_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 35 | val nat_divide_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 36 | val int_divide_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 37 | val nat_less_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 38 | val int_less_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 39 | val gcd_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 40 | val lcm_rel : n_ary_index | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
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: 
33982diff
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: 
36385diff
changeset | 57 | type kodkod_constrs = | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 58 |     {kk_all: decl list -> formula -> formula,
 | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 59 | kk_exist: decl list -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 60 | kk_formula_let: expr_assign list -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 61 | kk_formula_if: formula -> formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 62 | kk_or: formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 63 | kk_not: formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 64 | kk_iff: formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 65 | kk_implies: formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 66 | kk_and: formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 67 | kk_subset: rel_expr -> rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 68 | kk_rel_eq: rel_expr -> rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 69 | kk_no: rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 70 | kk_lone: rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 71 | kk_one: rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 72 | kk_some: rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
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: 
36385diff
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: 
36385diff
changeset | 75 | kk_union: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 76 | kk_difference: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 77 | kk_override: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 78 | kk_intersect: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 79 | kk_product: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 80 | kk_join: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 81 | kk_closure: rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 82 | kk_reflexive_closure: rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 83 | kk_comprehension: decl list -> formula -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 84 | kk_project: rel_expr -> int_expr list -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 85 | kk_project_seq: rel_expr -> int -> int -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 86 | kk_not3: rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 87 | kk_nat_less: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
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: 
33192diff
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: 
33192diff
changeset | 97 | open Nitpick_Util | 
| 33192 | 98 | |
| 36390 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 99 | type name_pool = | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 100 |   {rels: n_ary_index list,
 | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 101 | vars: n_ary_index list, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 102 | formula_reg: int, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
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: 
35284diff
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: 
33982diff
changeset | 139 | fun is_twos_complement_representable bits n = | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
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: 
33192diff
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: 
33192diff
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: 
33192diff
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: 
36385diff
changeset | 223 | type kodkod_constrs = | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 224 |   {kk_all: decl list -> formula -> formula,
 | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 225 | kk_exist: decl list -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 226 | kk_formula_let: expr_assign list -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 227 | kk_formula_if: formula -> formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 228 | kk_or: formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 229 | kk_not: formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 230 | kk_iff: formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 231 | kk_implies: formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 232 | kk_and: formula -> formula -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 233 | kk_subset: rel_expr -> rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 234 | kk_rel_eq: rel_expr -> rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 235 | kk_no: rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 236 | kk_lone: rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 237 | kk_one: rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 238 | kk_some: rel_expr -> formula, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
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: 
36385diff
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: 
36385diff
changeset | 241 | kk_union: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 242 | kk_difference: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 243 | kk_override: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 244 | kk_intersect: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 245 | kk_product: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 246 | kk_join: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 247 | kk_closure: rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 248 | kk_reflexive_closure: rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 249 | kk_comprehension: decl list -> formula -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 250 | kk_project: rel_expr -> int_expr list -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 251 | kk_project_seq: rel_expr -> int -> int -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 252 | kk_not3: rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
changeset | 253 | kk_nat_less: rel_expr -> rel_expr -> rel_expr, | 
| 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 blanchet parents: 
36385diff
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: 
35280diff
changeset | 332 | fun s_xone xone r = | 
| 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35280diff
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: 
35280diff
changeset | 334 | True | 
| 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35280diff
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: 
35280diff
changeset | 336 | 1 => xone r | 
| 
9edc2bd6d2bd
enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
 blanchet parents: 
35280diff
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: 
35280diff
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: 
35280diff
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: 
35280diff
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: 
33982diff
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: 
33982diff
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: 
33982diff
changeset | 359 | (Join (r11, Rel x), _) => | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
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: 
33982diff
changeset | 361 | | (_, Join (r21, Rel x)) => | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 362 | if x = not3_rel then s_rel_eq r21 (s_not3 r1) else raise SAME () | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 363 | | (RelIf (f, r11, r12), _) => | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 364 | if inline_rel_expr r2 then | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 365 | 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: 
33982diff
changeset | 366 | else | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 367 | raise SAME () | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 368 | | (_, RelIf (f, r21, r22)) => | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 369 | if inline_rel_expr r1 then | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 370 | 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: 
33982diff
changeset | 371 | else | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 372 | raise SAME () | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 373 | | (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: 
33982diff
changeset | 374 | | (Atom _, RelLet (bs, r2')) => s_formula_let bs (s_rel_eq r1 r2') | 
| 33192 | 375 | | _ => raise SAME ()) | 
| 376 | handle SAME () => | |
| 377 | case rel_expr_equal r1 r2 of | |
| 378 | SOME true => True | |
| 379 | | SOME false => False | |
| 380 | | NONE => | |
| 381 | case (r1, r2) of | |
| 382 | (_, RelIf (f, r21, r22)) => | |
| 383 | if inline_rel_expr r1 then | |
| 384 | s_formula_if f (s_rel_eq r1 r21) (s_rel_eq r1 r22) | |
| 385 | else | |
| 386 | RelEq (r1, r2) | |
| 387 | | (RelIf (f, r11, r12), _) => | |
| 388 | if inline_rel_expr r2 then | |
| 389 | s_formula_if f (s_rel_eq r11 r2) (s_rel_eq r12 r2) | |
| 390 | else | |
| 391 | RelEq (r1, r2) | |
| 34126 | 392 | | (_, None) => s_no r1 | 
| 393 | | (None, _) => s_no r2 | |
| 33192 | 394 | | _ => RelEq (r1, r2) | 
| 395 | fun s_subset (Atom j1) (Atom j2) = formula_for_bool (j1 = j2) | |
| 396 | | s_subset (Atom j) (AtomSeq (k, j0)) = | |
| 397 | formula_for_bool (j >= j0 andalso j < j0 + k) | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
34982diff
changeset | 398 | | s_subset (Union (r11, r12)) r2 = | 
| 33192 | 399 | s_and (s_subset r11 r2) (s_subset r12 r2) | 
| 400 | | s_subset r1 (r2 as Union (r21, r22)) = | |
| 401 | if is_one_rel_expr r1 then | |
| 402 | s_or (s_subset r1 r21) (s_subset r1 r22) | |
| 403 | else | |
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34126diff
changeset | 404 | 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: 
34126diff
changeset | 405 | r1 = r2 then | 
| 33192 | 406 | True | 
| 407 | else | |
| 408 | Subset (r1, r2) | |
| 409 | | s_subset r1 r2 = | |
| 410 | if r1 = r2 orelse is_none_product r1 then True | |
| 411 | else if is_none_product r2 then s_no r1 | |
| 412 | else if forall is_one_rel_expr [r1, r2] then s_rel_eq r1 r2 | |
| 413 | else Subset (r1, r2) | |
| 414 | ||
| 415 | fun s_rel_let [b as AssignRelReg (x', r')] (r as RelReg x) = | |
| 416 | if x = x' then r' else RelLet ([b], r) | |
| 417 | | s_rel_let bs r = RelLet (bs, r) | |
| 418 | ||
| 419 | fun s_rel_if f r1 r2 = | |
| 420 | (case (f, r1, r2) of | |
| 421 | (True, _, _) => r1 | |
| 422 | | (False, _, _) => r2 | |
| 423 | | (No r1', None, RelIf (One r2', r3', r4')) => | |
| 424 | if r1' = r2' andalso r2' = r3' then s_rel_if (Lone r1') r1' r4' | |
| 425 | else raise SAME () | |
| 426 | | _ => raise SAME ()) | |
| 427 | handle SAME () => if r1 = r2 then r1 else RelIf (f, r1, r2) | |
| 428 | ||
| 429 | fun s_union r1 (Union (r21, r22)) = s_union (s_union r1 r21) r22 | |
| 430 | | s_union r1 r2 = | |
| 431 | if is_none_product r1 then r2 | |
| 432 | else if is_none_product r2 then r1 | |
| 433 | else if r1 = r2 then r1 | |
| 434 | else if occurs_in_union r2 r1 then r1 | |
| 435 | else Union (r1, r2) | |
| 436 | fun s_difference r1 r2 = | |
| 437 | if is_none_product r1 orelse is_none_product r2 then r1 | |
| 438 | else if r1 = r2 then empty_n_ary_rel (arity_of_rel_expr r1) | |
| 439 | else Difference (r1, r2) | |
| 440 | fun s_override r1 r2 = | |
| 441 | if is_none_product r2 then r1 | |
| 442 | else if is_none_product r1 then r2 | |
| 443 | else Override (r1, r2) | |
| 444 | fun s_intersect r1 r2 = | |
| 445 | case rel_expr_intersects r1 r2 of | |
| 446 | SOME true => if r1 = r2 then r1 else Intersect (r1, r2) | |
| 447 | | SOME false => empty_n_ary_rel (arity_of_rel_expr r1) | |
| 448 | | NONE => if is_none_product r1 then r1 | |
| 449 | else if is_none_product r2 then r2 | |
| 450 | else Intersect (r1, r2) | |
| 451 | fun s_product r1 r2 = | |
| 452 | if is_none_product r1 then | |
| 453 | Product (r1, empty_n_ary_rel (arity_of_rel_expr r2)) | |
| 454 | else if is_none_product r2 then | |
| 455 | Product (empty_n_ary_rel (arity_of_rel_expr r1), r2) | |
| 456 | else | |
| 457 | Product (r1, r2) | |
| 458 | fun s_join r1 (Product (Product (r211, r212), r22)) = | |
| 459 | Product (s_join r1 (Product (r211, r212)), r22) | |
| 460 | | s_join (Product (r11, Product (r121, r122))) r2 = | |
| 461 | Product (r11, s_join (Product (r121, r122)) r2) | |
| 462 | | s_join None r = empty_n_ary_rel (arity_of_rel_expr r - 1) | |
| 463 | | s_join r None = empty_n_ary_rel (arity_of_rel_expr r - 1) | |
| 464 | | s_join (Product (None, None)) r = empty_n_ary_rel (arity_of_rel_expr r) | |
| 465 | | s_join r (Product (None, None)) = empty_n_ary_rel (arity_of_rel_expr r) | |
| 466 | | s_join Iden r2 = r2 | |
| 467 | | s_join r1 Iden = r1 | |
| 468 | | s_join (Product (r1, r2)) Univ = | |
| 469 | if arity_of_rel_expr r2 = 1 then r1 | |
| 470 | else Product (r1, s_join r2 Univ) | |
| 471 | | s_join Univ (Product (r1, r2)) = | |
| 472 | if arity_of_rel_expr r1 = 1 then r2 | |
| 473 | else Product (s_join Univ r1, r2) | |
| 474 | | s_join r1 (r2 as Product (r21, r22)) = | |
| 475 | if arity_of_rel_expr r1 = 1 then | |
| 476 | case rel_expr_intersects r1 r21 of | |
| 477 | SOME true => r22 | |
| 478 | | SOME false => empty_n_ary_rel (arity_of_rel_expr r2 - 1) | |
| 479 | | NONE => Join (r1, r2) | |
| 480 | else | |
| 481 | Join (r1, r2) | |
| 482 | | s_join (r1 as Product (r11, r12)) r2 = | |
| 483 | if arity_of_rel_expr r2 = 1 then | |
| 484 | case rel_expr_intersects r2 r12 of | |
| 485 | SOME true => r11 | |
| 486 | | SOME false => empty_n_ary_rel (arity_of_rel_expr r1 - 1) | |
| 487 | | NONE => Join (r1, r2) | |
| 488 | else | |
| 489 | Join (r1, r2) | |
| 490 | | s_join r1 (r2 as RelIf (f, r21, r22)) = | |
| 491 | if inline_rel_expr r1 then s_rel_if f (s_join r1 r21) (s_join r1 r22) | |
| 492 | else Join (r1, r2) | |
| 493 | | s_join (r1 as RelIf (f, r11, r12)) r2 = | |
| 494 | if inline_rel_expr r2 then s_rel_if f (s_join r11 r2) (s_join r12 r2) | |
| 495 | else Join (r1, r2) | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
34982diff
changeset | 496 | | 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: 
33982diff
changeset | 497 | if x = suc_rel then | 
| 33192 | 498 | let val n = to_nat j1 + 1 in | 
| 499 | if n < nat_card then from_nat n else None | |
| 500 | end | |
| 501 | else | |
| 502 | Join (r1, r2) | |
| 503 | | s_join r1 (r2 as Project (r21, Num k :: is)) = | |
| 504 | if k = arity_of_rel_expr r21 - 1 andalso arity_of_rel_expr r1 = 1 then | |
| 505 | s_project (s_join r21 r1) is | |
| 506 | else | |
| 507 | Join (r1, r2) | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
34982diff
changeset | 508 | | s_join r1 (Join (r21, r22 as Rel (x as (3, _)))) = | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 509 | ((if x = nat_add_rel then | 
| 33192 | 510 | case (r21, r1) of | 
| 511 | (Atom j1, Atom j2) => | |
| 512 | let val n = to_nat j1 + to_nat j2 in | |
| 513 | if n < nat_card then from_nat n else None | |
| 514 | end | |
| 515 | | (Atom j, r) => | |
| 516 | (case to_nat j of | |
| 517 | 0 => r | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 518 | | 1 => s_join r (Rel suc_rel) | 
| 33192 | 519 | | _ => raise SAME ()) | 
| 520 | | (r, Atom j) => | |
| 521 | (case to_nat j of | |
| 522 | 0 => r | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 523 | | 1 => s_join r (Rel suc_rel) | 
| 33192 | 524 | | _ => raise SAME ()) | 
| 525 | | _ => raise SAME () | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 526 | else if x = nat_subtract_rel then | 
| 33192 | 527 | case (r21, r1) of | 
| 33705 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 528 | (Atom j1, Atom j2) => from_nat (nat_minus (to_nat j1) (to_nat j2)) | 
| 33192 | 529 | | _ => raise SAME () | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 530 | else if x = nat_multiply_rel then | 
| 33192 | 531 | case (r21, r1) of | 
| 532 | (Atom j1, Atom j2) => | |
| 533 | let val n = to_nat j1 * to_nat j2 in | |
| 534 | if n < nat_card then from_nat n else None | |
| 535 | end | |
| 536 | | (Atom j, r) => | |
| 537 | (case to_nat j of 0 => Atom j | 1 => r | _ => raise SAME ()) | |
| 538 | | (r, Atom j) => | |
| 539 | (case to_nat j of 0 => Atom j | 1 => r | _ => raise SAME ()) | |
| 540 | | _ => raise SAME () | |
| 541 | else | |
| 542 | raise SAME ()) | |
| 543 | handle SAME () => List.foldr Join r22 [r1, r21]) | |
| 544 | | s_join r1 r2 = Join (r1, r2) | |
| 545 | ||
| 546 | fun s_closure Iden = Iden | |
| 547 | | s_closure r = if is_none_product r then r else Closure r | |
| 548 | fun s_reflexive_closure Iden = Iden | |
| 549 | | s_reflexive_closure r = | |
| 550 | if is_none_product r then Iden else ReflexiveClosure r | |
| 551 | ||
| 552 | fun s_comprehension ds False = empty_n_ary_rel (length ds) | |
| 553 | | s_comprehension ds True = fold1 s_product (map decl_one_set ds) | |
| 554 | | s_comprehension [d as DeclOne ((1, j1), r)] | |
| 555 | (f as RelEq (Var (1, j2), Atom j)) = | |
| 556 | if j1 = j2 andalso rel_expr_intersects (Atom j) r = SOME true then | |
| 557 | Atom j | |
| 558 | else | |
| 559 | Comprehension ([d], f) | |
| 560 | | s_comprehension ds f = Comprehension (ds, f) | |
| 561 | ||
| 562 | fun s_project_seq r = | |
| 563 | let | |
| 564 | fun aux arity r j0 n = | |
| 565 | if j0 = 0 andalso arity = n then | |
| 566 | r | |
| 567 | else case r of | |
| 568 | RelIf (f, r1, r2) => | |
| 569 | s_rel_if f (aux arity r1 j0 n) (aux arity r2 j0 n) | |
| 570 | | Product (r1, r2) => | |
| 571 | let | |
| 572 | val arity2 = arity_of_rel_expr r2 | |
| 573 | val arity1 = arity - arity2 | |
| 33705 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 574 | val n1 = Int.min (nat_minus arity1 j0, n) | 
| 33192 | 575 | val n2 = n - n1 | 
| 576 | fun one () = aux arity1 r1 j0 n1 | |
| 33705 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 577 | fun two () = aux arity2 r2 (nat_minus j0 arity1) n2 | 
| 33192 | 578 | in | 
| 579 | case (n1, n2) of | |
| 580 | (0, _) => s_rel_if (s_some r1) (two ()) (empty_n_ary_rel n2) | |
| 581 | | (_, 0) => s_rel_if (s_some r2) (one ()) (empty_n_ary_rel n1) | |
| 582 | | _ => s_product (one ()) (two ()) | |
| 583 | end | |
| 584 | | _ => s_project r (num_seq j0 n) | |
| 585 | in aux (arity_of_rel_expr r) r end | |
| 586 | ||
| 587 | 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: 
33982diff
changeset | 588 | | s_nat_less r1 r2 = fold s_join [r1, r2] (Rel nat_less_rel) | 
| 33192 | 589 | 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: 
33982diff
changeset | 590 | | s_int_less r1 r2 = fold s_join [r1, r2] (Rel int_less_rel) | 
| 33192 | 591 | |
| 592 | 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: 
33982diff
changeset | 593 | fun d_not3 r = Join (r, Rel not3_rel) | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
33982diff
changeset | 594 | 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: 
33982diff
changeset | 595 | fun d_int_less r1 r2 = List.foldl Join (Rel int_less_rel) [r1, r2] | 
| 33192 | 596 | in | 
| 597 | if optim then | |
| 598 |       {kk_all = s_all, kk_exist = s_exist, kk_formula_let = s_formula_let,
 | |
| 599 | kk_formula_if = s_formula_if, kk_or = s_or, kk_not = s_not, | |
| 600 | kk_iff = s_iff, kk_implies = s_implies, kk_and = s_and, | |
| 601 | kk_subset = s_subset, kk_rel_eq = s_rel_eq, kk_no = s_no, | |
| 602 | kk_lone = s_lone, kk_one = s_one, kk_some = s_some, | |
| 603 | kk_rel_let = s_rel_let, kk_rel_if = s_rel_if, kk_union = s_union, | |
| 604 | kk_difference = s_difference, kk_override = s_override, | |
| 605 | kk_intersect = s_intersect, kk_product = s_product, kk_join = s_join, | |
| 606 | kk_closure = s_closure, kk_reflexive_closure = s_reflexive_closure, | |
| 607 | kk_comprehension = s_comprehension, kk_project = s_project, | |
| 608 | kk_project_seq = s_project_seq, kk_not3 = s_not3, | |
| 609 | kk_nat_less = s_nat_less, kk_int_less = s_int_less} | |
| 610 | else | |
| 611 |       {kk_all = curry All, kk_exist = curry Exist,
 | |
| 612 | kk_formula_let = curry FormulaLet, kk_formula_if = curry3 FormulaIf, | |
| 613 | kk_or = curry Or,kk_not = Not, kk_iff = curry Iff, kk_implies = curry | |
| 614 | Implies, kk_and = curry And, kk_subset = curry Subset, kk_rel_eq = curry | |
| 615 | RelEq, kk_no = No, kk_lone = Lone, kk_one = One, kk_some = Some, | |
| 616 | kk_rel_let = curry RelLet, kk_rel_if = curry3 RelIf, kk_union = curry | |
| 617 | Union, kk_difference = curry Difference, kk_override = curry Override, | |
| 618 | kk_intersect = curry Intersect, kk_product = curry Product, | |
| 619 | kk_join = curry Join, kk_closure = Closure, | |
| 620 | kk_reflexive_closure = ReflexiveClosure, kk_comprehension = curry | |
| 621 | Comprehension, kk_project = curry Project, | |
| 622 | kk_project_seq = d_project_seq, kk_not3 = d_not3, | |
| 623 | kk_nat_less = d_nat_less, kk_int_less = d_int_less} | |
| 624 | end | |
| 625 | ||
| 626 | end; |