| author | nipkow | 
| Sat, 03 Dec 2011 21:25:34 +0100 | |
| changeset 45745 | 3a8bc5623410 | 
| parent 42436 | 0f60055d10fb | 
| child 46083 | efeaa79f021b | 
| permissions | -rw-r--r-- | 
| 33982 | 1 | (* Title: HOL/Tools/Nitpick/nitpick_kodkod.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 | Kodkod problem generator part of Kodkod. | |
| 6 | *) | |
| 7 | ||
| 8 | signature NITPICK_KODKOD = | |
| 9 | sig | |
| 35070 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 blanchet parents: 
34982diff
changeset | 10 | type hol_context = Nitpick_HOL.hol_context | 
| 38126 | 11 | type datatype_spec = Nitpick_Scope.datatype_spec | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 12 | type kodkod_constrs = Nitpick_Peephole.kodkod_constrs | 
| 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 13 | type nut = Nitpick_Nut.nut | 
| 33192 | 14 | |
| 15 | structure NameTable : TABLE | |
| 16 | ||
| 17 | val univ_card : | |
| 18 | int -> int -> int -> Kodkod.bound list -> Kodkod.formula -> int | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 19 | val check_bits : int -> Kodkod.formula -> unit | 
| 38182 | 20 | val check_arity : string -> int -> int -> unit | 
| 33192 | 21 | val kk_tuple : bool -> int -> int list -> Kodkod.tuple | 
| 22 | val tuple_set_from_atom_schema : (int * int) list -> Kodkod.tuple_set | |
| 23 | val sequential_int_bounds : int -> Kodkod.int_bound list | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 24 | val pow_of_two_int_bounds : int -> int -> Kodkod.int_bound list | 
| 38126 | 25 | val bounds_and_axioms_for_built_in_rels_in_formulas : | 
| 39345 | 26 | bool -> int -> int -> int -> int -> Kodkod.formula list | 
| 38126 | 27 | -> Kodkod.bound list * Kodkod.formula list | 
| 33192 | 28 | val bound_for_plain_rel : Proof.context -> bool -> nut -> Kodkod.bound | 
| 29 | val bound_for_sel_rel : | |
| 41995 | 30 | Proof.context -> bool -> (typ * (nut * int) list option) list | 
| 31 | -> datatype_spec list -> nut -> Kodkod.bound | |
| 33192 | 32 | val merge_bounds : Kodkod.bound list -> Kodkod.bound list | 
| 41802 | 33 | val kodkod_formula_from_nut : | 
| 34 | int Typtab.table -> kodkod_constrs -> nut -> Kodkod.formula | |
| 41995 | 35 | val needed_values_for_datatype : | 
| 36 | nut list -> int Typtab.table -> datatype_spec -> (nut * int) list option | |
| 33192 | 37 | val declarative_axiom_for_plain_rel : kodkod_constrs -> nut -> Kodkod.formula | 
| 38 | val declarative_axioms_for_datatypes : | |
| 41995 | 39 | hol_context -> bool -> nut list -> (typ * (nut * int) list option) list | 
| 40 | -> int -> int -> int Typtab.table -> kodkod_constrs -> nut NameTable.table | |
| 41 | -> datatype_spec list -> Kodkod.formula list | |
| 33192 | 42 | end; | 
| 43 | ||
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 44 | structure Nitpick_Kodkod : NITPICK_KODKOD = | 
| 33192 | 45 | struct | 
| 46 | ||
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 47 | open Nitpick_Util | 
| 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 48 | open Nitpick_HOL | 
| 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 49 | open Nitpick_Scope | 
| 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 50 | open Nitpick_Peephole | 
| 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 51 | open Nitpick_Rep | 
| 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 52 | open Nitpick_Nut | 
| 33192 | 53 | |
| 34126 | 54 | structure KK = Kodkod | 
| 55 | ||
| 38126 | 56 | fun pull x xs = x :: filter_out (curry (op =) x) xs | 
| 33192 | 57 | |
| 38196 | 58 | fun is_datatype_acyclic ({co = false, standard = true, deep = true, ...}
 | 
| 59 | : datatype_spec) = true | |
| 60 | | is_datatype_acyclic _ = false | |
| 33192 | 61 | |
| 34126 | 62 | fun flip_nums n = index_seq 1 n @ [0] |> map KK.Num | 
| 33192 | 63 | |
| 64 | fun univ_card nat_card int_card main_j0 bounds formula = | |
| 65 | let | |
| 66 | fun rel_expr_func r k = | |
| 67 | Int.max (k, case r of | |
| 34126 | 68 | KK.Atom j => j + 1 | 
| 69 | | KK.AtomSeq (k', j0) => j0 + k' | |
| 33192 | 70 | | _ => 0) | 
| 71 | fun tuple_func t k = | |
| 72 | case t of | |
| 34126 | 73 | KK.Tuple js => fold Integer.max (map (Integer.add 1) js) k | 
| 33192 | 74 | | _ => k | 
| 75 | fun tuple_set_func ts k = | |
| 34126 | 76 | Int.max (k, case ts of KK.TupleAtomSeq (k', j0) => j0 + k' | _ => 0) | 
| 33192 | 77 |     val expr_F = {formula_func = K I, rel_expr_func = rel_expr_func,
 | 
| 78 | int_expr_func = K I} | |
| 79 |     val tuple_F = {tuple_func = tuple_func, tuple_set_func = tuple_set_func}
 | |
| 34126 | 80 | val card = fold (KK.fold_bound expr_F tuple_F) bounds 1 | 
| 81 | |> KK.fold_formula expr_F formula | |
| 33192 | 82 | in Int.max (main_j0 + fold Integer.max [2, nat_card, int_card] 0, card) end | 
| 83 | ||
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 84 | fun check_bits bits formula = | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 85 | let | 
| 34126 | 86 | fun int_expr_func (KK.Num k) () = | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 87 | if is_twos_complement_representable bits k then | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 88 | () | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 89 | else | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 90 |           raise TOO_SMALL ("Nitpick_Kodkod.check_bits",
 | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 91 | "\"bits\" value " ^ string_of_int bits ^ | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 92 | " too small for problem") | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 93 | | int_expr_func _ () = () | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 94 |     val expr_F = {formula_func = K I, rel_expr_func = K I,
 | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 95 | int_expr_func = int_expr_func} | 
| 34126 | 96 | in KK.fold_formula expr_F formula () end | 
| 33192 | 97 | |
| 38182 | 98 | fun check_arity guilty_party univ_card n = | 
| 34126 | 99 | if n > KK.max_arity univ_card then | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 100 |     raise TOO_LARGE ("Nitpick_Kodkod.check_arity",
 | 
| 38182 | 101 | "arity " ^ string_of_int n ^ | 
| 102 | (if guilty_party = "" then | |
| 103 | "" | |
| 104 | else | |
| 105 | " of Kodkod relation associated with " ^ | |
| 41045 
2a41709f34c1
use heuristic to determine whether to keep or drop an existing "let" -- and drop all higher-order lets
 blanchet parents: 
39898diff
changeset | 106 | quote (original_name guilty_party)) ^ | 
| 38182 | 107 | " too large for universe of cardinality " ^ | 
| 108 | string_of_int univ_card) | |
| 33192 | 109 | else | 
| 110 | () | |
| 111 | ||
| 112 | fun kk_tuple debug univ_card js = | |
| 113 | if debug then | |
| 34126 | 114 | KK.Tuple js | 
| 33192 | 115 | else | 
| 34126 | 116 | KK.TupleIndex (length js, | 
| 117 | fold (fn j => fn accum => accum * univ_card + j) js 0) | |
| 33192 | 118 | |
| 42436 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 119 | (* This code is not just a silly optimization: It works around a limitation in | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 120 |    Kodkodi, whereby {} (i.e., KK.TupleProduct) is always given the cardinality
 | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 121 | of the bound in which it appears, resulting in Kodkod arity errors. *) | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 122 | fun tuple_product (ts as KK.TupleSet []) _ = ts | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 123 | | tuple_product _ (ts as KK.TupleSet []) = ts | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 124 | | tuple_product ts1 ts2 = KK.TupleProduct (ts1, ts2) | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 125 | |
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 126 | val tuple_set_from_atom_schema = fold1 tuple_product o map KK.TupleAtomSeq | 
| 33192 | 127 | val upper_bound_for_rep = tuple_set_from_atom_schema o atom_schema_of_rep | 
| 128 | ||
| 34126 | 129 | val single_atom = KK.TupleSet o single o KK.Tuple o single | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 130 | fun sequential_int_bounds n = [(NONE, map single_atom (index_seq 0 n))] | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 131 | fun pow_of_two_int_bounds bits j0 = | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 132 | let | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 133 | fun aux 0 _ _ = [] | 
| 35220 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 blanchet parents: 
35190diff
changeset | 134 | | aux 1 pow_of_two j = [(SOME (~ pow_of_two), [single_atom j])] | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 135 | | aux iter pow_of_two j = | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 136 | (SOME pow_of_two, [single_atom j]) :: | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 137 | aux (iter - 1) (2 * pow_of_two) (j + 1) | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 138 | in aux (bits + 1) 1 j0 end | 
| 33192 | 139 | |
| 38126 | 140 | fun built_in_rels_in_formulas formulas = | 
| 33192 | 141 | let | 
| 39345 | 142 | fun rel_expr_func (KK.Rel (x as (_, j))) = | 
| 38126 | 143 | (j < 0 andalso x <> unsigned_bit_word_sel_rel andalso | 
| 144 | x <> signed_bit_word_sel_rel) | |
| 145 | ? insert (op =) x | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 146 | | rel_expr_func _ = I | 
| 33192 | 147 |     val expr_F = {formula_func = K I, rel_expr_func = rel_expr_func,
 | 
| 148 | int_expr_func = K I} | |
| 38126 | 149 | in fold (KK.fold_formula expr_F) formulas [] end | 
| 33192 | 150 | |
| 151 | val max_table_size = 65536 | |
| 152 | ||
| 153 | fun check_table_size k = | |
| 154 | if k > max_table_size then | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 155 |     raise TOO_LARGE ("Nitpick_Kodkod.check_table_size",
 | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 156 |                      "precomputed table too large (" ^ string_of_int k ^ ")")
 | 
| 33192 | 157 | else | 
| 158 | () | |
| 159 | ||
| 160 | fun tabulate_func1 debug univ_card (k, j0) f = | |
| 161 | (check_table_size k; | |
| 162 | map_filter (fn j1 => let val j2 = f j1 in | |
| 163 | if j2 >= 0 then | |
| 164 | SOME (kk_tuple debug univ_card [j1 + j0, j2 + j0]) | |
| 165 | else | |
| 166 | NONE | |
| 167 | end) (index_seq 0 k)) | |
| 168 | fun tabulate_op2 debug univ_card (k, j0) res_j0 f = | |
| 169 | (check_table_size (k * k); | |
| 170 | map_filter (fn j => let | |
| 171 | val j1 = j div k | |
| 172 | val j2 = j - j1 * k | |
| 173 | val j3 = f (j1, j2) | |
| 174 | in | |
| 175 | if j3 >= 0 then | |
| 176 | SOME (kk_tuple debug univ_card | |
| 177 | [j1 + j0, j2 + j0, j3 + res_j0]) | |
| 178 | else | |
| 179 | NONE | |
| 180 | end) (index_seq 0 (k * k))) | |
| 181 | fun tabulate_op2_2 debug univ_card (k, j0) res_j0 f = | |
| 182 | (check_table_size (k * k); | |
| 183 | map_filter (fn j => let | |
| 184 | val j1 = j div k | |
| 185 | val j2 = j - j1 * k | |
| 186 | val (j3, j4) = f (j1, j2) | |
| 187 | in | |
| 188 | if j3 >= 0 andalso j4 >= 0 then | |
| 189 | SOME (kk_tuple debug univ_card | |
| 190 | [j1 + j0, j2 + j0, j3 + res_j0, | |
| 191 | j4 + res_j0]) | |
| 192 | else | |
| 193 | NONE | |
| 194 | end) (index_seq 0 (k * k))) | |
| 195 | fun tabulate_nat_op2 debug univ_card (k, j0) f = | |
| 196 | tabulate_op2 debug univ_card (k, j0) j0 (atom_for_nat (k, 0) o f) | |
| 197 | fun tabulate_int_op2 debug univ_card (k, j0) f = | |
| 198 | tabulate_op2 debug univ_card (k, j0) j0 | |
| 199 | (atom_for_int (k, 0) o f o pairself (int_for_atom (k, 0))) | |
| 200 | fun tabulate_int_op2_2 debug univ_card (k, j0) f = | |
| 201 | tabulate_op2_2 debug univ_card (k, j0) j0 | |
| 202 | (pairself (atom_for_int (k, 0)) o f | |
| 203 | o pairself (int_for_atom (k, 0))) | |
| 204 | ||
| 205 | fun isa_div (m, n) = m div n handle General.Div => 0 | |
| 206 | fun isa_mod (m, n) = m mod n handle General.Div => m | |
| 207 | fun isa_gcd (m, 0) = m | |
| 208 | | isa_gcd (m, n) = isa_gcd (n, isa_mod (m, n)) | |
| 209 | fun isa_lcm (m, n) = isa_div (m * n, isa_gcd (m, n)) | |
| 210 | val isa_zgcd = isa_gcd o pairself abs | |
| 211 | fun isa_norm_frac (m, n) = | |
| 212 | if n < 0 then isa_norm_frac (~m, ~n) | |
| 213 | else if m = 0 orelse n = 0 then (0, 1) | |
| 214 | else let val p = isa_zgcd (m, n) in (isa_div (m, p), isa_div (n, p)) end | |
| 215 | ||
| 39345 | 216 | fun tabulate_built_in_rel debug univ_card nat_card int_card j0 | 
| 38124 | 217 | (x as (n, _)) = | 
| 38182 | 218 | (check_arity "" univ_card n; | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 219 | if x = not3_rel then | 
| 33192 | 220 |      ("not3", tabulate_func1 debug univ_card (2, j0) (curry (op -) 1))
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 221 | else if x = suc_rel then | 
| 38126 | 222 |      ("suc", tabulate_func1 debug univ_card (univ_card - j0 - 1, j0)
 | 
| 223 | (Integer.add 1)) | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 224 | else if x = nat_add_rel then | 
| 33192 | 225 |      ("nat_add", tabulate_nat_op2 debug univ_card (nat_card, j0) (op +))
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 226 | else if x = int_add_rel then | 
| 33192 | 227 |      ("int_add", tabulate_int_op2 debug univ_card (int_card, j0) (op +))
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 228 | else if x = nat_subtract_rel then | 
| 33192 | 229 |      ("nat_subtract",
 | 
| 33705 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33631diff
changeset | 230 | tabulate_op2 debug univ_card (nat_card, j0) j0 (uncurry nat_minus)) | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 231 | else if x = int_subtract_rel then | 
| 33192 | 232 |      ("int_subtract", tabulate_int_op2 debug univ_card (int_card, j0) (op -))
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 233 | else if x = nat_multiply_rel then | 
| 33192 | 234 |      ("nat_multiply", tabulate_nat_op2 debug univ_card (nat_card, j0) (op * ))
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 235 | else if x = int_multiply_rel then | 
| 33192 | 236 |      ("int_multiply", tabulate_int_op2 debug univ_card (int_card, j0) (op * ))
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 237 | else if x = nat_divide_rel then | 
| 33192 | 238 |      ("nat_divide", tabulate_nat_op2 debug univ_card (nat_card, j0) isa_div)
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 239 | else if x = int_divide_rel then | 
| 33192 | 240 |      ("int_divide", tabulate_int_op2 debug univ_card (int_card, j0) isa_div)
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 241 | else if x = nat_less_rel then | 
| 33192 | 242 |      ("nat_less", tabulate_nat_op2 debug univ_card (nat_card, j0)
 | 
| 35385 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 blanchet parents: 
35312diff
changeset | 243 | (int_from_bool o op <)) | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 244 | else if x = int_less_rel then | 
| 33192 | 245 |      ("int_less", tabulate_int_op2 debug univ_card (int_card, j0)
 | 
| 35385 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 blanchet parents: 
35312diff
changeset | 246 | (int_from_bool o op <)) | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 247 | else if x = gcd_rel then | 
| 33192 | 248 |      ("gcd", tabulate_nat_op2 debug univ_card (nat_card, j0) isa_gcd)
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 249 | else if x = lcm_rel then | 
| 33192 | 250 |      ("lcm", tabulate_nat_op2 debug univ_card (nat_card, j0) isa_lcm)
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 251 | else if x = norm_frac_rel then | 
| 33192 | 252 |      ("norm_frac", tabulate_int_op2_2 debug univ_card (int_card, j0)
 | 
| 253 | isa_norm_frac) | |
| 254 | else | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 255 |      raise ARG ("Nitpick_Kodkod.tabulate_built_in_rel", "unknown relation"))
 | 
| 33192 | 256 | |
| 39345 | 257 | fun bound_for_built_in_rel debug univ_card nat_card int_card main_j0 | 
| 38126 | 258 | (x as (n, j)) = | 
| 259 | if n = 2 andalso j <= suc_rels_base then | |
| 260 | let val (y as (k, j0), tabulate) = atom_seq_for_suc_rel x in | |
| 261 | ([(x, "suc")], | |
| 262 | if tabulate then | |
| 263 | [KK.TupleSet (tabulate_func1 debug univ_card (k - 1, j0) | |
| 264 | (Integer.add 1))] | |
| 265 | else | |
| 266 | [KK.TupleSet [], tuple_set_from_atom_schema [y, y]]) | |
| 267 | end | |
| 268 | else | |
| 269 | let | |
| 39345 | 270 | val (nick, ts) = tabulate_built_in_rel debug univ_card nat_card int_card | 
| 271 | main_j0 x | |
| 38126 | 272 | in ([(x, nick)], [KK.TupleSet ts]) end | 
| 33192 | 273 | |
| 38126 | 274 | fun axiom_for_built_in_rel (x as (n, j)) = | 
| 275 | if n = 2 andalso j <= suc_rels_base then | |
| 276 | let val (y as (k, j0), tabulate) = atom_seq_for_suc_rel x in | |
| 38160 | 277 | if tabulate then | 
| 38126 | 278 | NONE | 
| 38160 | 279 | else if k < 2 then | 
| 280 | SOME (KK.No (KK.Rel x)) | |
| 38126 | 281 | else | 
| 282 | SOME (KK.TotalOrdering (x, KK.AtomSeq y, KK.Atom j0, KK.Atom (j0 + 1))) | |
| 283 | end | |
| 284 | else | |
| 285 | NONE | |
| 39345 | 286 | fun bounds_and_axioms_for_built_in_rels_in_formulas debug univ_card nat_card | 
| 38126 | 287 | int_card main_j0 formulas = | 
| 288 | let val rels = built_in_rels_in_formulas formulas in | |
| 39345 | 289 | (map (bound_for_built_in_rel debug univ_card nat_card int_card main_j0) | 
| 38126 | 290 | rels, | 
| 291 | map_filter axiom_for_built_in_rel rels) | |
| 39316 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 blanchet parents: 
38212diff
changeset | 292 | end | 
| 33192 | 293 | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 294 | fun bound_comment ctxt debug nick T R = | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 295 | short_name nick ^ | 
| 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 | 296 | (if debug then " :: " ^ unyxml (Syntax.string_of_typ ctxt T) else "") ^ | 
| 
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 | 297 | " : " ^ string_for_rep R | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 298 | |
| 33192 | 299 | fun bound_for_plain_rel ctxt debug (u as FreeRel (x, T, R, nick)) = | 
| 300 | ([(x, bound_comment ctxt debug nick T R)], | |
| 301 |      if nick = @{const_name bisim_iterator_max} then
 | |
| 302 | case R of | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 303 | Atom (k, j0) => [single_atom (k - 1 + j0)] | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 304 |        | _ => raise NUT ("Nitpick_Kodkod.bound_for_plain_rel", [u])
 | 
| 33192 | 305 | else | 
| 34126 | 306 | [KK.TupleSet [], upper_bound_for_rep R]) | 
| 33192 | 307 | | bound_for_plain_rel _ _ u = | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 308 |     raise NUT ("Nitpick_Kodkod.bound_for_plain_rel", [u])
 | 
| 33192 | 309 | |
| 41997 | 310 | fun is_datatype_nat_like ({typ, constrs, ...} : datatype_spec) =
 | 
| 42000 | 311 | case constrs of | 
| 312 | [_, _] => | |
| 313 | (case constrs |> map (snd o #const) |> List.partition is_fun_type of | |
| 314 | ([Type (_, Ts1)], [T2]) => forall (curry (op =) typ) (T2 :: Ts1) | |
| 315 | | _ => false) | |
| 41997 | 316 | | _ => false | 
| 317 | ||
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 318 | fun needed_values need_vals T = | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 319 | AList.lookup (op =) need_vals T |> the_default NONE |> these | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 320 | |
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 321 | fun all_values_are_needed need_vals ({typ, card, ...} : datatype_spec) =
 | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 322 | length (needed_values need_vals typ) = card | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 323 | |
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 324 | fun is_sel_of_constr x (Construct (sel_us, _, _, _), _) = | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 325 | exists (fn FreeRel (x', _, _, _) => x = x' | _ => false) sel_us | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 326 | | is_sel_of_constr _ _ = false | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 327 | |
| 41995 | 328 | fun bound_for_sel_rel ctxt debug need_vals dtypes | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 329 |         (FreeRel (x, T as Type (@{type_name fun}, [T1, T2]),
 | 
| 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 330 | R as Func (Atom (_, j0), R2), nick)) = | 
| 33192 | 331 | let | 
| 41995 | 332 | val constr_s = original_name nick | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 333 |       val {delta, epsilon, exclusive, explicit_max, ...} =
 | 
| 41995 | 334 | constr_spec dtypes (constr_s, T1) | 
| 335 |       val dtype as {card, ...} = datatype_spec dtypes T1 |> the
 | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 336 | val T1_need_vals = needed_values need_vals T1 | 
| 33192 | 337 | in | 
| 338 | ([(x, bound_comment ctxt debug nick T R)], | |
| 41995 | 339 | let | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 340 | val discr = (R2 = Formula Neut) | 
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 341 | val complete_need_vals = (length T1_need_vals = card) | 
| 41995 | 342 | val (my_need_vals, other_need_vals) = | 
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 343 | T1_need_vals |> List.partition (is_sel_of_constr x) | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 344 | fun atom_seq_for_self_rec j = | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 345 | if is_datatype_nat_like dtype then (1, j + j0 - 1) else (j, j0) | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 346 | fun exact_bound_for_perhaps_needy_atom j = | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 347 | case AList.find (op =) my_need_vals j of | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 348 | [constr_u] => | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 349 | let | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 350 | val n = sel_no_from_name nick | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 351 | val arg_u = | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 352 | case constr_u of | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 353 | Construct (_, _, _, arg_us) => nth arg_us n | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 354 | | _ => raise Fail "expected \"Construct\"" | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 355 | val T2_need_vals = needed_values need_vals T2 | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 356 | in | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 357 | case AList.lookup (op =) T2_need_vals arg_u of | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 358 | SOME arg_j => SOME (KK.TupleAtomSeq (1, arg_j)) | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 359 | | _ => NONE | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 360 | end | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 361 | | _ => NONE | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 362 | fun n_fold_tuple_union [] = KK.TupleSet [] | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 363 | | n_fold_tuple_union (ts :: tss) = | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 364 | fold (curry (KK.TupleUnion o swap)) tss ts | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 365 | fun tuple_perhaps_needy_atom upper j = | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 366 | single_atom j | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 367 | |> not discr | 
| 42436 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 368 | ? (fn ts => tuple_product ts | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 369 | (case exact_bound_for_perhaps_needy_atom j of | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 370 | SOME ts => ts | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 371 | | NONE => if upper then upper_bound_for_rep R2 | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 372 | else KK.TupleSet [])) | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 373 | fun bound_tuples upper = | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 374 | if null T1_need_vals then | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 375 | if upper then | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 376 | KK.TupleAtomSeq (epsilon - delta, delta + j0) | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 377 | |> not discr | 
| 42436 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 378 | ? (fn ts => tuple_product ts (upper_bound_for_rep R2)) | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 379 | else | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 380 | KK.TupleSet [] | 
| 41996 | 381 | else | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 382 | (if complete_need_vals then | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 383 | my_need_vals |> map snd | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 384 | else | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 385 | index_seq (delta + j0) (epsilon - delta) | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 386 | |> filter_out (member (op = o apsnd snd) other_need_vals)) | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 387 | |> map (tuple_perhaps_needy_atom upper) | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 388 | |> n_fold_tuple_union | 
| 41995 | 389 | in | 
| 390 | if explicit_max = 0 orelse | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 391 | (complete_need_vals andalso null my_need_vals) then | 
| 41995 | 392 | [KK.TupleSet []] | 
| 393 | else | |
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 394 | if discr then | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 395 | [bound_tuples true] | 
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 396 | |> not (exclusive orelse all_values_are_needed need_vals dtype) | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 397 | ? cons (KK.TupleSet []) | 
| 33192 | 398 | else | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 399 | [bound_tuples false, | 
| 35178 | 400 | if T1 = T2 andalso epsilon > delta andalso | 
| 41995 | 401 | is_datatype_acyclic dtype then | 
| 35072 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 blanchet parents: 
35070diff
changeset | 402 | index_seq delta (epsilon - delta) | 
| 42436 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 403 | |> map (fn j => tuple_product (KK.TupleSet [KK.Tuple [j + j0]]) | 
| 
0f60055d10fb
worked around Kodkodi limitation with parsing {}, where it cannot easily deduce its arity -- without this workaround, Kodkod sometimes generates arity errors in conjunction with the "need" option since change 614ff13dc5d2
 blanchet parents: 
42001diff
changeset | 404 | (KK.TupleAtomSeq (atom_seq_for_self_rec j))) | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 405 | |> n_fold_tuple_union | 
| 35072 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 blanchet parents: 
35070diff
changeset | 406 | else | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 407 | bound_tuples true] | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 408 | |> distinct (op =) | 
| 33192 | 409 | end) | 
| 410 | end | |
| 41995 | 411 | | bound_for_sel_rel _ _ _ _ u = | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 412 |     raise NUT ("Nitpick_Kodkod.bound_for_sel_rel", [u])
 | 
| 33192 | 413 | |
| 414 | fun merge_bounds bs = | |
| 415 | let | |
| 416 | fun arity (zs, _) = fst (fst (hd zs)) | |
| 417 | fun add_bound ds b [] = List.revAppend (ds, [b]) | |
| 418 | | add_bound ds b (c :: cs) = | |
| 419 | if arity b = arity c andalso snd b = snd c then | |
| 420 | List.revAppend (ds, (fst c @ fst b, snd c) :: cs) | |
| 421 | else | |
| 422 | add_bound (c :: ds) b cs | |
| 423 | in fold (add_bound []) bs [] end | |
| 424 | ||
| 34126 | 425 | fun unary_var_seq j0 n = map (curry KK.Var 1) (index_seq j0 n) | 
| 33192 | 426 | |
| 34126 | 427 | val singleton_from_combination = foldl1 KK.Product o map KK.Atom | 
| 33192 | 428 | fun all_singletons_for_rep R = | 
| 429 | if is_lone_rep R then | |
| 430 | all_combinations_for_rep R |> map singleton_from_combination | |
| 431 | else | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 432 |     raise REP ("Nitpick_Kodkod.all_singletons_for_rep", [R])
 | 
| 33192 | 433 | |
| 34126 | 434 | fun unpack_products (KK.Product (r1, r2)) = | 
| 33192 | 435 | unpack_products r1 @ unpack_products r2 | 
| 436 | | unpack_products r = [r] | |
| 34126 | 437 | fun unpack_joins (KK.Join (r1, r2)) = unpack_joins r1 @ unpack_joins r2 | 
| 33192 | 438 | | unpack_joins r = [r] | 
| 439 | ||
| 440 | val empty_rel_for_rep = empty_n_ary_rel o arity_of_rep | |
| 441 | fun full_rel_for_rep R = | |
| 442 | case atom_schema_of_rep R of | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 443 |     [] => raise REP ("Nitpick_Kodkod.full_rel_for_rep", [R])
 | 
| 34126 | 444 | | schema => foldl1 KK.Product (map KK.AtomSeq schema) | 
| 33192 | 445 | |
| 446 | fun decls_for_atom_schema j0 schema = | |
| 34126 | 447 | map2 (fn j => fn x => KK.DeclOne ((1, j), KK.AtomSeq x)) | 
| 33192 | 448 | (index_seq j0 (length schema)) schema | 
| 449 | ||
| 450 | fun d_n_ary_function ({kk_all, kk_join, kk_lone, kk_one, ...} : kodkod_constrs)
 | |
| 451 | R r = | |
| 452 | let val body_R = body_rep R in | |
| 453 | if is_lone_rep body_R then | |
| 454 | let | |
| 455 | val binder_schema = atom_schema_of_reps (binder_reps R) | |
| 456 | val body_schema = atom_schema_of_rep body_R | |
| 457 | val one = is_one_rep body_R | |
| 34126 | 458 | val opt_x = case r of KK.Rel x => SOME x | _ => NONE | 
| 33192 | 459 | in | 
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 460 | if opt_x <> NONE andalso length binder_schema = 1 andalso | 
| 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 461 | length body_schema = 1 then | 
| 34126 | 462 | (if one then KK.Function else KK.Functional) | 
| 463 | (the opt_x, KK.AtomSeq (hd binder_schema), | |
| 464 | KK.AtomSeq (hd body_schema)) | |
| 33192 | 465 | else | 
| 466 | let | |
| 467 | val decls = decls_for_atom_schema ~1 binder_schema | |
| 468 | val vars = unary_var_seq ~1 (length binder_schema) | |
| 469 | val kk_xone = if one then kk_one else kk_lone | |
| 470 | in kk_all decls (kk_xone (fold kk_join vars r)) end | |
| 471 | end | |
| 472 | else | |
| 34126 | 473 | KK.True | 
| 33192 | 474 | end | 
| 34126 | 475 | fun kk_n_ary_function kk R (r as KK.Rel x) = | 
| 33192 | 476 | if not (is_opt_rep R) then | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 477 | if x = suc_rel then | 
| 34126 | 478 | KK.False | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 479 | else if x = nat_add_rel then | 
| 33192 | 480 | formula_for_bool (card_of_rep (body_rep R) = 1) | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 481 | else if x = nat_multiply_rel then | 
| 33192 | 482 | formula_for_bool (card_of_rep (body_rep R) <= 2) | 
| 483 | else | |
| 484 | d_n_ary_function kk R r | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 485 | else if x = nat_subtract_rel then | 
| 34126 | 486 | KK.True | 
| 33192 | 487 | else | 
| 488 | d_n_ary_function kk R r | |
| 489 | | kk_n_ary_function kk R r = d_n_ary_function kk R r | |
| 490 | ||
| 34126 | 491 | fun kk_disjoint_sets _ [] = KK.True | 
| 33192 | 492 |   | kk_disjoint_sets (kk as {kk_and, kk_no, kk_intersect, ...} : kodkod_constrs)
 | 
| 493 | (r :: rs) = | |
| 494 | fold (kk_and o kk_no o kk_intersect r) rs (kk_disjoint_sets kk rs) | |
| 495 | ||
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 496 | fun basic_rel_rel_let j ({kk_rel_let, ...} : kodkod_constrs) f r =
 | 
| 33192 | 497 | if inline_rel_expr r then | 
| 498 | f r | |
| 499 | else | |
| 34126 | 500 | let val x = (KK.arity_of_rel_expr r, j) in | 
| 501 | kk_rel_let [KK.AssignRelReg (x, r)] (f (KK.RelReg x)) | |
| 33192 | 502 | end | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 503 | val single_rel_rel_let = basic_rel_rel_let 0 | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 504 | fun double_rel_rel_let kk f r1 r2 = | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 505 | single_rel_rel_let kk (fn r1 => basic_rel_rel_let 1 kk (f r1) r2) r1 | 
| 35386 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 506 | fun triple_rel_rel_let kk f r1 r2 r3 = | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 507 | double_rel_rel_let kk | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 508 | (fn r1 => fn r2 => basic_rel_rel_let 2 kk (f r1 r2) r3) r1 r2 | 
| 33192 | 509 | |
| 510 | fun atom_from_formula ({kk_rel_if, ...} : kodkod_constrs) j0 f =
 | |
| 34126 | 511 | kk_rel_if f (KK.Atom (j0 + 1)) (KK.Atom j0) | 
| 33192 | 512 | fun rel_expr_from_formula kk R f = | 
| 513 | case unopt_rep R of | |
| 514 | Atom (2, j0) => atom_from_formula kk j0 f | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 515 |   | _ => raise REP ("Nitpick_Kodkod.rel_expr_from_formula", [R])
 | 
| 33192 | 516 | |
| 517 | fun unpack_vect_in_chunks ({kk_project_seq, ...} : kodkod_constrs) chunk_arity
 | |
| 518 | num_chunks r = | |
| 519 | List.tabulate (num_chunks, fn j => kk_project_seq r (j * chunk_arity) | |
| 520 | chunk_arity) | |
| 521 | ||
| 522 | fun kk_n_fold_join | |
| 523 |         (kk as {kk_intersect, kk_product, kk_join, kk_project_seq, ...}) one R1
 | |
| 524 | res_R r1 r2 = | |
| 525 | case arity_of_rep R1 of | |
| 526 | 1 => kk_join r1 r2 | |
| 527 | | arity1 => | |
| 38164 | 528 | let val unpacked_rs1 = unpack_products r1 in | 
| 33192 | 529 | if one andalso length unpacked_rs1 = arity1 then | 
| 530 | fold kk_join unpacked_rs1 r2 | |
| 38164 | 531 | else if one andalso inline_rel_expr r1 then | 
| 532 | fold kk_join (unpack_vect_in_chunks kk 1 arity1 r1) r2 | |
| 33192 | 533 | else | 
| 534 | kk_project_seq | |
| 535 | (kk_intersect (kk_product r1 (full_rel_for_rep res_R)) r2) | |
| 536 | arity1 (arity_of_rep res_R) | |
| 537 | end | |
| 538 | ||
| 539 | fun kk_case_switch (kk as {kk_union, kk_product, ...}) R1 R2 r rs1 rs2 =
 | |
| 540 | if rs1 = rs2 then r | |
| 541 | else kk_n_fold_join kk true R1 R2 r (fold1 kk_union (map2 kk_product rs1 rs2)) | |
| 542 | ||
| 543 | val lone_rep_fallback_max_card = 4096 | |
| 544 | val some_j0 = 0 | |
| 545 | ||
| 546 | fun lone_rep_fallback kk new_R old_R r = | |
| 547 | if old_R = new_R then | |
| 548 | r | |
| 549 | else | |
| 550 | let val card = card_of_rep old_R in | |
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 551 | if is_lone_rep old_R andalso is_lone_rep new_R andalso | 
| 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 552 | card = card_of_rep new_R then | 
| 33192 | 553 | if card >= lone_rep_fallback_max_card then | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 554 |           raise TOO_LARGE ("Nitpick_Kodkod.lone_rep_fallback",
 | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 555 |                            "too high cardinality (" ^ string_of_int card ^ ")")
 | 
| 33192 | 556 | else | 
| 557 | kk_case_switch kk old_R new_R r (all_singletons_for_rep old_R) | |
| 558 | (all_singletons_for_rep new_R) | |
| 559 | else | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 560 |         raise REP ("Nitpick_Kodkod.lone_rep_fallback", [old_R, new_R])
 | 
| 33192 | 561 | end | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 562 | and atom_from_rel_expr kk x old_R r = | 
| 33192 | 563 | case old_R of | 
| 564 | Func (R1, R2) => | |
| 565 | let | |
| 566 | val dom_card = card_of_rep R1 | |
| 567 | val R2' = case R2 of Atom _ => R2 | _ => Atom (card_of_rep R2, some_j0) | |
| 568 | in | |
| 569 | atom_from_rel_expr kk x (Vect (dom_card, R2')) | |
| 570 | (vect_from_rel_expr kk dom_card R2' old_R r) | |
| 571 | end | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 572 |   | Opt _ => raise REP ("Nitpick_Kodkod.atom_from_rel_expr", [old_R])
 | 
| 33192 | 573 | | _ => lone_rep_fallback kk (Atom x) old_R r | 
| 574 | and struct_from_rel_expr kk Rs old_R r = | |
| 575 | case old_R of | |
| 576 | Atom _ => lone_rep_fallback kk (Struct Rs) old_R r | |
| 577 | | Struct Rs' => | |
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 578 | if Rs' = Rs then | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 579 | r | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 580 | else if map card_of_rep Rs' = map card_of_rep Rs then | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 581 | let | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 582 | val old_arities = map arity_of_rep Rs' | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 583 | val old_offsets = offset_list old_arities | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 584 | val old_rs = map2 (#kk_project_seq kk r) old_offsets old_arities | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 585 | in | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 586 | fold1 (#kk_product kk) | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 587 | (map3 (rel_expr_from_rel_expr kk) Rs Rs' old_rs) | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 588 | end | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 589 | else | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 590 | lone_rep_fallback kk (Struct Rs) old_R r | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 591 |   | _ => raise REP ("Nitpick_Kodkod.struct_from_rel_expr", [old_R])
 | 
| 33192 | 592 | and vect_from_rel_expr kk k R old_R r = | 
| 593 | case old_R of | |
| 594 | Atom _ => lone_rep_fallback kk (Vect (k, R)) old_R r | |
| 595 | | Vect (k', R') => | |
| 596 | if k = k' andalso R = R' then r | |
| 597 | else lone_rep_fallback kk (Vect (k, R)) old_R r | |
| 598 | | Func (R1, Formula Neut) => | |
| 599 | if k = card_of_rep R1 then | |
| 600 | fold1 (#kk_product kk) | |
| 601 | (map (fn arg_r => | |
| 602 | rel_expr_from_formula kk R (#kk_subset kk arg_r r)) | |
| 603 | (all_singletons_for_rep R1)) | |
| 604 | else | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 605 |       raise REP ("Nitpick_Kodkod.vect_from_rel_expr", [old_R])
 | 
| 33192 | 606 | | Func (R1, R2) => | 
| 607 | fold1 (#kk_product kk) | |
| 608 | (map (fn arg_r => | |
| 609 | rel_expr_from_rel_expr kk R R2 | |
| 610 | (kk_n_fold_join kk true R1 R2 arg_r r)) | |
| 611 | (all_singletons_for_rep R1)) | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 612 |   | _ => raise REP ("Nitpick_Kodkod.vect_from_rel_expr", [old_R])
 | 
| 33192 | 613 | and func_from_no_opt_rel_expr kk R1 R2 (Atom x) r = | 
| 614 | let | |
| 615 | val dom_card = card_of_rep R1 | |
| 616 | val R2' = case R2 of Atom _ => R2 | _ => Atom (card_of_rep R2, some_j0) | |
| 617 | in | |
| 618 | func_from_no_opt_rel_expr kk R1 R2 (Vect (dom_card, R2')) | |
| 619 | (vect_from_rel_expr kk dom_card R2' (Atom x) r) | |
| 620 | end | |
| 621 | | func_from_no_opt_rel_expr kk R1 (Formula Neut) old_R r = | |
| 622 | (case old_R of | |
| 623 | Vect (k, Atom (2, j0)) => | |
| 624 | let | |
| 625 | val args_rs = all_singletons_for_rep R1 | |
| 626 | val vals_rs = unpack_vect_in_chunks kk 1 k r | |
| 627 | fun empty_or_singleton_set_for arg_r val_r = | |
| 34126 | 628 | #kk_join kk val_r (#kk_product kk (KK.Atom (j0 + 1)) arg_r) | 
| 33192 | 629 | in | 
| 630 | fold1 (#kk_union kk) (map2 empty_or_singleton_set_for args_rs vals_rs) | |
| 631 | end | |
| 632 | | Func (R1', Formula Neut) => | |
| 633 | if R1 = R1' then | |
| 634 | r | |
| 635 | else | |
| 636 | let | |
| 637 | val schema = atom_schema_of_rep R1 | |
| 638 | val r1 = fold1 (#kk_product kk) (unary_var_seq ~1 (length schema)) | |
| 639 | |> rel_expr_from_rel_expr kk R1' R1 | |
| 33582 
bdf98e327f0b
fixed soundness bug in Nitpick related to sets of sets;
 blanchet parents: 
33571diff
changeset | 640 | val kk_xeq = (if is_one_rep R1' then #kk_subset else #kk_rel_eq) kk | 
| 33192 | 641 | in | 
| 33582 
bdf98e327f0b
fixed soundness bug in Nitpick related to sets of sets;
 blanchet parents: 
33571diff
changeset | 642 | #kk_comprehension kk (decls_for_atom_schema ~1 schema) (kk_xeq r1 r) | 
| 33192 | 643 | end | 
| 644 | | Func (R1', Atom (2, j0)) => | |
| 645 | func_from_no_opt_rel_expr kk R1 (Formula Neut) | |
| 34126 | 646 | (Func (R1', Formula Neut)) (#kk_join kk r (KK.Atom (j0 + 1))) | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 647 |      | _ => raise REP ("Nitpick_Kodkod.func_from_no_opt_rel_expr",
 | 
| 33192 | 648 | [old_R, Func (R1, Formula Neut)])) | 
| 649 | | func_from_no_opt_rel_expr kk R1 R2 old_R r = | |
| 650 | case old_R of | |
| 651 | Vect (k, R) => | |
| 652 | let | |
| 653 | val args_rs = all_singletons_for_rep R1 | |
| 654 | val vals_rs = unpack_vect_in_chunks kk (arity_of_rep R) k r | |
| 655 | |> map (rel_expr_from_rel_expr kk R2 R) | |
| 656 | in fold1 (#kk_union kk) (map2 (#kk_product kk) args_rs vals_rs) end | |
| 657 | | Func (R1', Formula Neut) => | |
| 658 | (case R2 of | |
| 659 | Atom (x as (2, j0)) => | |
| 660 | let val schema = atom_schema_of_rep R1 in | |
| 661 | if length schema = 1 then | |
| 34126 | 662 | #kk_override kk (#kk_product kk (KK.AtomSeq (hd schema)) | 
| 663 | (KK.Atom j0)) | |
| 664 | (#kk_product kk r (KK.Atom (j0 + 1))) | |
| 33192 | 665 | else | 
| 666 | let | |
| 667 | val r1 = fold1 (#kk_product kk) (unary_var_seq ~1 (length schema)) | |
| 668 | |> rel_expr_from_rel_expr kk R1' R1 | |
| 34126 | 669 | val r2 = KK.Var (1, ~(length schema) - 1) | 
| 33192 | 670 | val r3 = atom_from_formula kk j0 (#kk_subset kk r1 r) | 
| 671 | in | |
| 672 | #kk_comprehension kk (decls_for_atom_schema ~1 (schema @ [x])) | |
| 33582 
bdf98e327f0b
fixed soundness bug in Nitpick related to sets of sets;
 blanchet parents: 
33571diff
changeset | 673 | (#kk_subset kk r2 r3) | 
| 33192 | 674 | end | 
| 675 | end | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 676 |          | _ => raise REP ("Nitpick_Kodkod.func_from_no_opt_rel_expr",
 | 
| 33192 | 677 | [old_R, Func (R1, R2)])) | 
| 678 | | Func (R1', R2') => | |
| 679 | if R1 = R1' andalso R2 = R2' then | |
| 680 | r | |
| 681 | else | |
| 682 | let | |
| 683 | val dom_schema = atom_schema_of_rep R1 | |
| 684 | val ran_schema = atom_schema_of_rep R2 | |
| 685 | val dom_prod = fold1 (#kk_product kk) | |
| 686 | (unary_var_seq ~1 (length dom_schema)) | |
| 687 | |> rel_expr_from_rel_expr kk R1' R1 | |
| 688 | val ran_prod = fold1 (#kk_product kk) | |
| 689 | (unary_var_seq (~(length dom_schema) - 1) | |
| 690 | (length ran_schema)) | |
| 691 | |> rel_expr_from_rel_expr kk R2' R2 | |
| 692 | val app = kk_n_fold_join kk true R1' R2' dom_prod r | |
| 33582 
bdf98e327f0b
fixed soundness bug in Nitpick related to sets of sets;
 blanchet parents: 
33571diff
changeset | 693 | val kk_xeq = (if is_one_rep R2' then #kk_subset else #kk_rel_eq) kk | 
| 33192 | 694 | in | 
| 695 | #kk_comprehension kk (decls_for_atom_schema ~1 | |
| 696 | (dom_schema @ ran_schema)) | |
| 33582 
bdf98e327f0b
fixed soundness bug in Nitpick related to sets of sets;
 blanchet parents: 
33571diff
changeset | 697 | (kk_xeq ran_prod app) | 
| 33192 | 698 | end | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 699 |     | _ => raise REP ("Nitpick_Kodkod.func_from_no_opt_rel_expr",
 | 
| 33192 | 700 | [old_R, Func (R1, R2)]) | 
| 701 | and rel_expr_from_rel_expr kk new_R old_R r = | |
| 702 | let | |
| 703 | val unopt_old_R = unopt_rep old_R | |
| 704 | val unopt_new_R = unopt_rep new_R | |
| 705 | in | |
| 706 | if unopt_old_R <> old_R andalso unopt_new_R = new_R then | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 707 |       raise REP ("Nitpick_Kodkod.rel_expr_from_rel_expr", [old_R, new_R])
 | 
| 33192 | 708 | else if unopt_new_R = unopt_old_R then | 
| 709 | r | |
| 710 | else | |
| 711 | (case unopt_new_R of | |
| 712 | Atom x => atom_from_rel_expr kk x | |
| 713 | | Struct Rs => struct_from_rel_expr kk Rs | |
| 714 | | Vect (k, R') => vect_from_rel_expr kk k R' | |
| 715 | | Func (R1, R2) => func_from_no_opt_rel_expr kk R1 R2 | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 716 |        | _ => raise REP ("Nitpick_Kodkod.rel_expr_from_rel_expr",
 | 
| 33192 | 717 | [old_R, new_R])) | 
| 718 | unopt_old_R r | |
| 719 | end | |
| 720 | and rel_expr_to_func kk R1 R2 = rel_expr_from_rel_expr kk (Func (R1, R2)) | |
| 721 | ||
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 722 | fun bit_set_from_atom ({kk_join, ...} : kodkod_constrs) T r =
 | 
| 34126 | 723 |   kk_join r (KK.Rel (if T = @{typ "unsigned_bit word"} then
 | 
| 724 | unsigned_bit_word_sel_rel | |
| 725 | else | |
| 726 | signed_bit_word_sel_rel)) | |
| 727 | val int_expr_from_atom = KK.SetSum ooo bit_set_from_atom | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 728 | fun atom_from_int_expr (kk as {kk_rel_eq, kk_comprehension, ...}
 | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 729 | : kodkod_constrs) T R i = | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 730 | kk_comprehension (decls_for_atom_schema ~1 (atom_schema_of_rep R)) | 
| 34126 | 731 | (kk_rel_eq (bit_set_from_atom kk T (KK.Var (1, ~1))) | 
| 732 | (KK.Bits i)) | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 733 | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 734 | fun kodkod_formula_from_nut ofs | 
| 33192 | 735 |         (kk as {kk_all, kk_exist, kk_formula_let, kk_formula_if, kk_or, kk_not,
 | 
| 35072 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 blanchet parents: 
35070diff
changeset | 736 | kk_iff, kk_implies, kk_and, kk_subset, kk_rel_eq, kk_no, | 
| 39456 | 737 | kk_lone, kk_some, kk_rel_let, kk_rel_if, kk_union, | 
| 35072 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 blanchet parents: 
35070diff
changeset | 738 | kk_difference, kk_intersect, kk_product, kk_join, kk_closure, | 
| 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 blanchet parents: 
35070diff
changeset | 739 | kk_comprehension, kk_project, kk_project_seq, kk_not3, | 
| 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 blanchet parents: 
35070diff
changeset | 740 | kk_nat_less, kk_int_less, ...}) u = | 
| 33192 | 741 | let | 
| 742 | val main_j0 = offset_of_type ofs bool_T | |
| 743 | val bool_j0 = main_j0 | |
| 744 | val bool_atom_R = Atom (2, main_j0) | |
| 34126 | 745 | val false_atom = KK.Atom bool_j0 | 
| 746 | val true_atom = KK.Atom (bool_j0 + 1) | |
| 33192 | 747 | fun formula_from_opt_atom polar j0 r = | 
| 748 | case polar of | |
| 34126 | 749 | Neg => kk_not (kk_rel_eq r (KK.Atom j0)) | 
| 750 | | _ => kk_rel_eq r (KK.Atom (j0 + 1)) | |
| 33192 | 751 | val formula_from_atom = formula_from_opt_atom Pos | 
| 752 | val unpack_formulas = | |
| 753 | map (formula_from_atom bool_j0) oo unpack_vect_in_chunks kk 1 | |
| 754 | fun kk_vect_set_bool_op connective k r1 r2 = | |
| 755 | fold1 kk_and (map2 connective (unpack_formulas k r1) | |
| 756 | (unpack_formulas k r2)) | |
| 757 | fun to_f u = | |
| 758 | case rep_of u of | |
| 759 | Formula polar => | |
| 760 | (case u of | |
| 34126 | 761 | Cst (False, _, _) => KK.False | 
| 762 | | Cst (True, _, _) => KK.True | |
| 33854 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 763 | | Op1 (Not, _, _, u1) => | 
| 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 764 | kk_not (to_f_with_polarity (flip_polarity polar) u1) | 
| 33192 | 765 | | Op1 (Finite, _, _, u1) => | 
| 766 | let val opt1 = is_opt_rep (rep_of u1) in | |
| 767 | case polar of | |
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 768 | Neut => | 
| 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 769 |                if opt1 then raise NUT ("Nitpick_Kodkod.to_f (Finite)", [u])
 | 
| 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 770 | else KK.True | 
| 33192 | 771 | | Pos => formula_for_bool (not opt1) | 
| 34126 | 772 | | Neg => KK.True | 
| 33192 | 773 | end | 
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 774 | | Op1 (IsUnknown, _, _, u1) => kk_no (to_r u1) | 
| 33192 | 775 | | Op1 (Cast, _, _, u1) => to_f_with_polarity polar u1 | 
| 33854 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 776 | | Op2 (All, _, _, u1, u2) => | 
| 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 777 | kk_all (untuple to_decl u1) (to_f_with_polarity polar u2) | 
| 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 778 | | Op2 (Exist, _, _, u1, u2) => | 
| 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 779 | kk_exist (untuple to_decl u1) (to_f_with_polarity polar u2) | 
| 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 780 | | Op2 (Or, _, _, u1, u2) => | 
| 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 781 | kk_or (to_f_with_polarity polar u1) (to_f_with_polarity polar u2) | 
| 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 782 | | Op2 (And, _, _, u1, u2) => | 
| 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 783 | kk_and (to_f_with_polarity polar u1) (to_f_with_polarity polar u2) | 
| 33192 | 784 | | Op2 (Less, T, Formula polar, u1, u2) => | 
| 785 | formula_from_opt_atom polar bool_j0 | |
| 786 | (to_r (Op2 (Less, T, Opt bool_atom_R, u1, u2))) | |
| 787 | | Op2 (Subset, _, _, u1, u2) => | |
| 788 | let | |
| 789 | val dom_T = domain_type (type_of u1) | |
| 790 | val R1 = rep_of u1 | |
| 791 | val R2 = rep_of u2 | |
| 792 | val (dom_R, ran_R) = | |
| 793 | case min_rep R1 R2 of | |
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 794 | Func Rp => Rp | 
| 33192 | 795 | | R => (Atom (card_of_domain_from_rep 2 R, | 
| 796 | offset_of_type ofs dom_T), | |
| 797 | if is_opt_rep R then Opt bool_atom_R else Formula Neut) | |
| 798 | val set_R = Func (dom_R, ran_R) | |
| 799 | in | |
| 800 | if not (is_opt_rep ran_R) then | |
| 801 | to_set_bool_op kk_implies kk_subset u1 u2 | |
| 802 | else if polar = Neut then | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 803 |                raise NUT ("Nitpick_Kodkod.to_f (Subset)", [u])
 | 
| 33192 | 804 | else | 
| 805 | let | |
| 33886 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 806 | (* FIXME: merge with similar code below *) | 
| 33192 | 807 | fun set_to_r widen u = | 
| 808 | if widen then | |
| 809 | kk_difference (full_rel_for_rep dom_R) | |
| 810 | (kk_join (to_rep set_R u) false_atom) | |
| 811 | else | |
| 812 | kk_join (to_rep set_R u) true_atom | |
| 813 | val widen1 = (polar = Pos andalso is_opt_rep R1) | |
| 814 | val widen2 = (polar = Neg andalso is_opt_rep R2) | |
| 815 | in kk_subset (set_to_r widen1 u1) (set_to_r widen2 u2) end | |
| 816 | end | |
| 817 | | Op2 (DefEq, _, _, u1, u2) => | |
| 818 | (case min_rep (rep_of u1) (rep_of u2) of | |
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 819 | Formula polar => | 
| 33192 | 820 | kk_iff (to_f_with_polarity polar u1) (to_f_with_polarity polar u2) | 
| 821 | | min_R => | |
| 822 | let | |
| 823 | fun args (Op2 (Apply, _, _, u1, u2)) = u2 :: args u1 | |
| 824 | | args (Tuple (_, _, us)) = us | |
| 825 | | args _ = [] | |
| 826 | val opt_arg_us = filter (is_opt_rep o rep_of) (args u1) | |
| 827 | in | |
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 828 | if null opt_arg_us orelse not (is_Opt min_R) orelse | 
| 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 829 | is_eval_name u1 then | 
| 33192 | 830 | fold (kk_or o (kk_no o to_r)) opt_arg_us | 
| 831 | (kk_rel_eq (to_rep min_R u1) (to_rep min_R u2)) | |
| 832 | else | |
| 34121 
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
 blanchet parents: 
33982diff
changeset | 833 | kk_subset (to_rep min_R u1) (to_rep min_R u2) | 
| 33192 | 834 | end) | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 835 | | Op2 (Eq, _, _, u1, u2) => | 
| 33192 | 836 | (case min_rep (rep_of u1) (rep_of u2) of | 
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 837 | Formula polar => | 
| 33192 | 838 | kk_iff (to_f_with_polarity polar u1) (to_f_with_polarity polar u2) | 
| 839 | | min_R => | |
| 840 | if is_opt_rep min_R then | |
| 841 | if polar = Neut then | |
| 842 | (* continuation of hackish optimization *) | |
| 843 | kk_rel_eq (to_rep min_R u1) (to_rep min_R u2) | |
| 844 | else if is_Cst Unrep u1 then | |
| 845 | to_could_be_unrep (polar = Neg) u2 | |
| 846 | else if is_Cst Unrep u2 then | |
| 847 | to_could_be_unrep (polar = Neg) u1 | |
| 848 | else | |
| 849 | let | |
| 850 | val r1 = to_rep min_R u1 | |
| 851 | val r2 = to_rep min_R u2 | |
| 852 | val both_opt = forall (is_opt_rep o rep_of) [u1, u2] | |
| 853 | in | |
| 854 | (if polar = Pos then | |
| 855 | if not both_opt then | |
| 856 | kk_rel_eq r1 r2 | |
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 857 | else if is_lone_rep min_R andalso | 
| 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 858 | arity_of_rep min_R = 1 then | 
| 33192 | 859 | kk_some (kk_intersect r1 r2) | 
| 860 | else | |
| 861 | raise SAME () | |
| 862 | else | |
| 863 | if is_lone_rep min_R then | |
| 864 | if arity_of_rep min_R = 1 then | |
| 35072 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 blanchet parents: 
35070diff
changeset | 865 | kk_lone (kk_union r1 r2) | 
| 33192 | 866 | else if not both_opt then | 
| 867 | (r1, r2) |> is_opt_rep (rep_of u2) ? swap | |
| 34121 
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
 blanchet parents: 
33982diff
changeset | 868 | |-> kk_subset | 
| 33192 | 869 | else | 
| 870 | raise SAME () | |
| 871 | else | |
| 872 | raise SAME ()) | |
| 873 | handle SAME () => | |
| 874 | formula_from_opt_atom polar bool_j0 | |
| 875 | (to_guard [u1, u2] bool_atom_R | |
| 876 | (rel_expr_from_formula kk bool_atom_R | |
| 877 | (kk_rel_eq r1 r2))) | |
| 878 | end | |
| 879 | else | |
| 880 | let | |
| 881 | val r1 = to_rep min_R u1 | |
| 882 | val r2 = to_rep min_R u2 | |
| 883 | in | |
| 884 | if is_one_rep min_R then | |
| 885 | let | |
| 886 | val rs1 = unpack_products r1 | |
| 887 | val rs2 = unpack_products r2 | |
| 888 | in | |
| 34936 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 889 | if length rs1 = length rs2 andalso | 
| 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 890 | map KK.arity_of_rel_expr rs1 | 
| 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 blanchet parents: 
34288diff
changeset | 891 | = map KK.arity_of_rel_expr rs2 then | 
| 33192 | 892 | fold1 kk_and (map2 kk_subset rs1 rs2) | 
| 893 | else | |
| 894 | kk_subset r1 r2 | |
| 895 | end | |
| 896 | else | |
| 897 | kk_rel_eq r1 r2 | |
| 898 | end) | |
| 899 | | Op2 (Apply, T, _, u1, u2) => | |
| 900 | (case (polar, rep_of u1) of | |
| 901 | (Neg, Func (R, Formula Neut)) => kk_subset (to_opt R u2) (to_r u1) | |
| 902 | | _ => | |
| 903 | to_f_with_polarity polar | |
| 904 | (Op2 (Apply, T, Opt (Atom (2, offset_of_type ofs T)), u1, u2))) | |
| 905 | | Op3 (Let, _, _, u1, u2, u3) => | |
| 33854 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 906 | kk_formula_let [to_expr_assign u1 u2] (to_f_with_polarity polar u3) | 
| 33192 | 907 | | Op3 (If, _, _, u1, u2, u3) => | 
| 33854 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 908 | kk_formula_if (to_f u1) (to_f_with_polarity polar u2) | 
| 
26a4cb3a7eae
fixed a Kodkod generation exception in Nitpick, reported by a Karlsruhe user
 blanchet parents: 
33744diff
changeset | 909 | (to_f_with_polarity polar u3) | 
| 34126 | 910 | | FormulaReg (j, _, _) => KK.FormulaReg j | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 911 |          | _ => raise NUT ("Nitpick_Kodkod.to_f", [u]))
 | 
| 33192 | 912 | | Atom (2, j0) => formula_from_atom j0 (to_r u) | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 913 |       | _ => raise NUT ("Nitpick_Kodkod.to_f", [u])
 | 
| 33192 | 914 | and to_f_with_polarity polar u = | 
| 915 | case rep_of u of | |
| 916 | Formula _ => to_f u | |
| 917 | | Atom (2, j0) => formula_from_atom j0 (to_r u) | |
| 918 | | Opt (Atom (2, j0)) => formula_from_opt_atom polar j0 (to_r u) | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 919 |       | _ => raise NUT ("Nitpick_Kodkod.to_f_with_polarity", [u])
 | 
| 33192 | 920 | and to_r u = | 
| 921 | case u of | |
| 922 | Cst (False, _, Atom _) => false_atom | |
| 923 | | Cst (True, _, Atom _) => true_atom | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 924 | | Cst (Iden, _, Func (Struct [R1, R2], Formula Neut)) => | 
| 33192 | 925 | if R1 = R2 andalso arity_of_rep R1 = 1 then | 
| 34126 | 926 | kk_intersect KK.Iden (kk_product (full_rel_for_rep R1) KK.Univ) | 
| 33192 | 927 | else | 
| 928 | let | |
| 929 | val schema1 = atom_schema_of_rep R1 | |
| 930 | val schema2 = atom_schema_of_rep R2 | |
| 931 | val arity1 = length schema1 | |
| 932 | val arity2 = length schema2 | |
| 933 | val r1 = fold1 kk_product (unary_var_seq 0 arity1) | |
| 934 | val r2 = fold1 kk_product (unary_var_seq arity1 arity2) | |
| 935 | val min_R = min_rep R1 R2 | |
| 936 | in | |
| 937 | kk_comprehension | |
| 938 | (decls_for_atom_schema 0 (schema1 @ schema2)) | |
| 939 | (kk_rel_eq (rel_expr_from_rel_expr kk min_R R1 r1) | |
| 940 | (rel_expr_from_rel_expr kk min_R R2 r2)) | |
| 941 | end | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 942 | | Cst (Iden, _, Func (Atom (1, j0), Formula Neut)) => KK.Atom j0 | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 943 |       | Cst (Iden, T as Type (@{type_name fun}, [T1, _]), R as Func (R1, _)) =>
 | 
| 33192 | 944 | to_rep R (Cst (Iden, T, Func (one_rep ofs T1 R1, Formula Neut))) | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 945 | | Cst (Num j, T, R) => | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 946 | if is_word_type T then | 
| 34126 | 947 | atom_from_int_expr kk T R (KK.Num j) | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 948 | else if T = int_T then | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 949 | case atom_for_int (card_of_rep R, offset_of_type ofs int_T) j of | 
| 34126 | 950 | ~1 => if is_opt_rep R then KK.None | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 951 |                   else raise NUT ("Nitpick_Kodkod.to_r (Num)", [u])
 | 
| 34126 | 952 | | j' => KK.Atom j' | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 953 | else | 
| 34126 | 954 | if j < card_of_rep R then KK.Atom (j + offset_of_type ofs T) | 
| 955 | else if is_opt_rep R then KK.None | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 956 |           else raise NUT ("Nitpick_Kodkod.to_r (Num)", [u])
 | 
| 33192 | 957 | | Cst (Unknown, _, R) => empty_rel_for_rep R | 
| 958 | | Cst (Unrep, _, R) => empty_rel_for_rep R | |
| 34126 | 959 |       | Cst (Suc, T as @{typ "unsigned_bit word => unsigned_bit word"}, R) =>
 | 
| 960 | to_bit_word_unary_op T R (curry KK.Add (KK.Num 1)) | |
| 961 |       | Cst (Suc, @{typ "nat => nat"}, Func (Atom x, _)) =>
 | |
| 962 | kk_intersect (KK.Rel suc_rel) (kk_product KK.Univ (KK.AtomSeq x)) | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 963 | | Cst (Suc, _, Func (Atom _, _)) => KK.Rel suc_rel | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 964 |       | Cst (Add, Type (_, [@{typ nat}, _]), _) => KK.Rel nat_add_rel
 | 
| 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 965 |       | Cst (Add, Type (_, [@{typ int}, _]), _) => KK.Rel int_add_rel
 | 
| 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 966 |       | Cst (Add, T as Type (_, [@{typ "unsigned_bit word"}, _]), R) =>
 | 
| 34126 | 967 | to_bit_word_binary_op T R NONE (SOME (curry KK.Add)) | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 968 |       | Cst (Add, T as Type (_, [@{typ "signed_bit word"}, _]), R) =>
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 969 | to_bit_word_binary_op T R | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 970 | (SOME (fn i1 => fn i2 => fn i3 => | 
| 34126 | 971 | kk_implies (KK.LE (KK.Num 0, KK.BitXor (i1, i2))) | 
| 972 | (KK.LE (KK.Num 0, KK.BitXor (i2, i3))))) | |
| 973 | (SOME (curry KK.Add)) | |
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 974 |       | Cst (Subtract, Type (_, [@{typ nat}, _]), _) =>
 | 
| 34126 | 975 | KK.Rel nat_subtract_rel | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 976 |       | Cst (Subtract, Type (_, [@{typ int}, _]), _) =>
 | 
| 34126 | 977 | KK.Rel int_subtract_rel | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 978 |       | Cst (Subtract, T as Type (_, [@{typ "unsigned_bit word"}, _]), R) =>
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 979 | to_bit_word_binary_op T R NONE | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 980 | (SOME (fn i1 => fn i2 => | 
| 34126 | 981 | KK.IntIf (KK.LE (i1, i2), KK.Num 0, KK.Sub (i1, i2)))) | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 982 |       | Cst (Subtract, T as Type (_, [@{typ "signed_bit word"}, _]), R) =>
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 983 | to_bit_word_binary_op T R | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 984 | (SOME (fn i1 => fn i2 => fn i3 => | 
| 34126 | 985 | kk_implies (KK.LT (KK.BitXor (i1, i2), KK.Num 0)) | 
| 986 | (KK.LT (KK.BitXor (i2, i3), KK.Num 0)))) | |
| 987 | (SOME (curry KK.Sub)) | |
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 988 |       | Cst (Multiply, Type (_, [@{typ nat}, _]), _) =>
 | 
| 34126 | 989 | KK.Rel nat_multiply_rel | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 990 |       | Cst (Multiply, Type (_, [@{typ int}, _]), _) =>
 | 
| 34126 | 991 | KK.Rel int_multiply_rel | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 992 | | Cst (Multiply, | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 993 |              T as Type (_, [Type (@{type_name word}, [bit_T]), _]), R) =>
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 994 | to_bit_word_binary_op T R | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 995 | (SOME (fn i1 => fn i2 => fn i3 => | 
| 34126 | 996 | kk_or (KK.IntEq (i2, KK.Num 0)) | 
| 997 | (KK.IntEq (KK.Div (i3, i2), i1) | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 998 |                        |> bit_T = @{typ signed_bit}
 | 
| 34126 | 999 | ? kk_and (KK.LE (KK.Num 0, | 
| 1000 | foldl1 KK.BitAnd [i1, i2, i3]))))) | |
| 1001 | (SOME (curry KK.Mult)) | |
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 1002 |       | Cst (Divide, Type (_, [@{typ nat}, _]), _) => KK.Rel nat_divide_rel
 | 
| 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 1003 |       | Cst (Divide, Type (_, [@{typ int}, _]), _) => KK.Rel int_divide_rel
 | 
| 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 1004 |       | Cst (Divide, T as Type (_, [@{typ "unsigned_bit word"}, _]), R) =>
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1005 | to_bit_word_binary_op T R NONE | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1006 | (SOME (fn i1 => fn i2 => | 
| 34126 | 1007 | KK.IntIf (KK.IntEq (i2, KK.Num 0), | 
| 1008 | KK.Num 0, KK.Div (i1, i2)))) | |
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 1009 |       | Cst (Divide, T as Type (_, [@{typ "signed_bit word"}, _]), R) =>
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1010 | to_bit_word_binary_op T R | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1011 | (SOME (fn i1 => fn i2 => fn i3 => | 
| 34126 | 1012 | KK.LE (KK.Num 0, foldl1 KK.BitAnd [i1, i2, i3]))) | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1013 | (SOME (fn i1 => fn i2 => | 
| 34126 | 1014 | KK.IntIf (kk_and (KK.LT (i1, KK.Num 0)) | 
| 1015 | (KK.LT (KK.Num 0, i2)), | |
| 1016 | KK.Sub (KK.Div (KK.Add (i1, KK.Num 1), i2), KK.Num 1), | |
| 1017 | KK.IntIf (kk_and (KK.LT (KK.Num 0, i1)) | |
| 1018 | (KK.LT (i2, KK.Num 0)), | |
| 1019 | KK.Sub (KK.Div (KK.Sub (i1, KK.Num 1), i2), KK.Num 1), | |
| 1020 | KK.IntIf (KK.IntEq (i2, KK.Num 0), | |
| 1021 | KK.Num 0, KK.Div (i1, i2)))))) | |
| 1022 | | Cst (Gcd, _, _) => KK.Rel gcd_rel | |
| 1023 | | Cst (Lcm, _, _) => KK.Rel lcm_rel | |
| 1024 | | Cst (Fracs, _, Func (Atom (1, _), _)) => KK.None | |
| 33192 | 1025 | | Cst (Fracs, _, Func (Struct _, _)) => | 
| 34126 | 1026 | kk_project_seq (KK.Rel norm_frac_rel) 2 2 | 
| 1027 | | Cst (NormFrac, _, _) => KK.Rel norm_frac_rel | |
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 1028 |       | Cst (NatToInt, Type (_, [@{typ nat}, _]), Func (Atom _, Atom _)) =>
 | 
| 34126 | 1029 | KK.Iden | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 1030 |       | Cst (NatToInt, Type (_, [@{typ nat}, _]),
 | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1031 | Func (Atom (_, nat_j0), Opt (Atom (int_k, int_j0)))) => | 
| 33192 | 1032 | if nat_j0 = int_j0 then | 
| 34126 | 1033 | kk_intersect KK.Iden | 
| 1034 | (kk_product (KK.AtomSeq (max_int_for_card int_k + 1, nat_j0)) | |
| 1035 | KK.Univ) | |
| 33192 | 1036 | else | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1037 |           raise BAD ("Nitpick_Kodkod.to_r (NatToInt)", "\"nat_j0 <> int_j0\"")
 | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 1038 |       | Cst (NatToInt, T as Type (_, [@{typ "unsigned_bit word"}, _]), R) =>
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1039 | to_bit_word_unary_op T R I | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 1040 |       | Cst (IntToNat, Type (_, [@{typ int}, _]),
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1041 | Func (Atom (int_k, int_j0), nat_R)) => | 
| 33192 | 1042 | let | 
| 1043 | val abs_card = max_int_for_card int_k + 1 | |
| 1044 | val (nat_k, nat_j0) = the_single (atom_schema_of_rep nat_R) | |
| 1045 | val overlap = Int.min (nat_k, abs_card) | |
| 1046 | in | |
| 1047 | if nat_j0 = int_j0 then | |
| 34126 | 1048 | kk_union (kk_product (KK.AtomSeq (int_k - abs_card, | 
| 1049 | int_j0 + abs_card)) | |
| 1050 | (KK.Atom nat_j0)) | |
| 1051 | (kk_intersect KK.Iden | |
| 1052 | (kk_product (KK.AtomSeq (overlap, int_j0)) KK.Univ)) | |
| 33192 | 1053 | else | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1054 |             raise BAD ("Nitpick_Kodkod.to_r (IntToNat)", "\"nat_j0 <> int_j0\"")
 | 
| 33192 | 1055 | end | 
| 35665 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 blanchet parents: 
35407diff
changeset | 1056 |       | Cst (IntToNat, T as Type (_, [@{typ "signed_bit word"}, _]), R) =>
 | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1057 | to_bit_word_unary_op T R | 
| 34126 | 1058 | (fn i => KK.IntIf (KK.LE (i, KK.Num 0), KK.Num 0, i)) | 
| 33192 | 1059 | | Op1 (Not, _, R, u1) => kk_not3 (to_rep R u1) | 
| 34126 | 1060 | | Op1 (Finite, _, Opt (Atom _), _) => KK.None | 
| 33192 | 1061 | | Op1 (Converse, T, R, u1) => | 
| 1062 | let | |
| 1063 | val (b_T, a_T) = HOLogic.dest_prodT (domain_type T) | |
| 1064 | val (b_R, a_R) = | |
| 1065 | case R of | |
| 1066 | Func (Struct [R1, R2], _) => (R1, R2) | |
| 1067 | | Func (R1, _) => | |
| 1068 | if card_of_rep R1 <> 1 then | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1069 |                 raise REP ("Nitpick_Kodkod.to_r (Converse)", [R])
 | 
| 33192 | 1070 | else | 
| 1071 | pairself (Atom o pair 1 o offset_of_type ofs) (b_T, a_T) | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1072 |             | _ => raise REP ("Nitpick_Kodkod.to_r (Converse)", [R])
 | 
| 33192 | 1073 | val body_R = body_rep R | 
| 1074 | val a_arity = arity_of_rep a_R | |
| 1075 | val b_arity = arity_of_rep b_R | |
| 1076 | val ab_arity = a_arity + b_arity | |
| 1077 | val body_arity = arity_of_rep body_R | |
| 1078 | in | |
| 1079 | kk_project (to_rep (Func (Struct [a_R, b_R], body_R)) u1) | |
| 34126 | 1080 | (map KK.Num (index_seq a_arity b_arity @ | 
| 1081 | index_seq 0 a_arity @ | |
| 1082 | index_seq ab_arity body_arity)) | |
| 33192 | 1083 | |> rel_expr_from_rel_expr kk R (Func (Struct [b_R, a_R], body_R)) | 
| 1084 | end | |
| 1085 | | Op1 (Closure, _, R, u1) => | |
| 1086 | if is_opt_rep R then | |
| 1087 | let | |
| 1088 | val T1 = type_of u1 | |
| 1089 | val R' = rep_to_binary_rel_rep ofs T1 (unopt_rep (rep_of u1)) | |
| 1090 | val R'' = opt_rep ofs T1 R' | |
| 1091 | in | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1092 | single_rel_rel_let kk | 
| 33192 | 1093 | (fn r => | 
| 1094 | let | |
| 1095 | val true_r = kk_closure (kk_join r true_atom) | |
| 1096 | val full_r = full_rel_for_rep R' | |
| 1097 | val false_r = kk_difference full_r | |
| 1098 | (kk_closure (kk_difference full_r | |
| 1099 | (kk_join r false_atom))) | |
| 1100 | in | |
| 1101 | rel_expr_from_rel_expr kk R R'' | |
| 1102 | (kk_union (kk_product true_r true_atom) | |
| 1103 | (kk_product false_r false_atom)) | |
| 1104 | end) (to_rep R'' u1) | |
| 1105 | end | |
| 1106 | else | |
| 1107 | let val R' = rep_to_binary_rel_rep ofs (type_of u1) (rep_of u1) in | |
| 1108 | rel_expr_from_rel_expr kk R R' (kk_closure (to_rep R' u1)) | |
| 1109 | end | |
| 1110 | | Op1 (SingletonSet, _, Func (R1, Opt _), Cst (Unrep, _, _)) => | |
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 1111 | kk_product (full_rel_for_rep R1) false_atom | 
| 33192 | 1112 | | Op1 (SingletonSet, _, R, u1) => | 
| 1113 | (case R of | |
| 1114 | Func (R1, Formula Neut) => to_rep R1 u1 | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1115 | | Func (R1, Opt _) => | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1116 | single_rel_rel_let kk | 
| 33192 | 1117 | (fn r => kk_rel_if (kk_no r) (empty_rel_for_rep R) | 
| 1118 | (rel_expr_to_func kk R1 bool_atom_R | |
| 1119 | (Func (R1, Formula Neut)) r)) | |
| 1120 | (to_opt R1 u1) | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1121 |          | _ => raise NUT ("Nitpick_Kodkod.to_r (SingletonSet)", [u]))
 | 
| 35671 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 blanchet parents: 
35665diff
changeset | 1122 | | Op1 (SafeThe, _, R, u1) => | 
| 33192 | 1123 | if is_opt_rep R then | 
| 1124 | kk_join (to_rep (Func (unopt_rep R, Opt bool_atom_R)) u1) true_atom | |
| 1125 | else | |
| 1126 | to_rep (Func (R, Formula Neut)) u1 | |
| 39345 | 1127 | | Op1 (First, _, R, u1) => to_nth_pair_sel 0 R u1 | 
| 1128 | | Op1 (Second, _, R, u1) => to_nth_pair_sel 1 R u1 | |
| 33192 | 1129 | | Op1 (Cast, _, R, u1) => | 
| 1130 | ((case rep_of u1 of | |
| 1131 | Formula _ => | |
| 1132 | (case unopt_rep R of | |
| 1133 | Atom (2, j0) => atom_from_formula kk j0 (to_f u1) | |
| 1134 | | _ => raise SAME ()) | |
| 1135 | | _ => raise SAME ()) | |
| 1136 | handle SAME () => rel_expr_from_rel_expr kk R (rep_of u1) (to_r u1)) | |
| 1137 | | Op2 (All, T, R as Opt _, u1, u2) => | |
| 1138 | to_r (Op1 (Not, T, R, | |
| 1139 | Op2 (Exist, T, R, u1, Op1 (Not, T, rep_of u2, u2)))) | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1140 | | Op2 (Exist, _, Opt _, u1, u2) => | 
| 33192 | 1141 | let val rs1 = untuple to_decl u1 in | 
| 1142 | if not (is_opt_rep (rep_of u2)) then | |
| 34126 | 1143 | kk_rel_if (kk_exist rs1 (to_f u2)) true_atom KK.None | 
| 33192 | 1144 | else | 
| 1145 | let val r2 = to_r u2 in | |
| 1146 | kk_union (kk_rel_if (kk_exist rs1 (kk_rel_eq r2 true_atom)) | |
| 34126 | 1147 | true_atom KK.None) | 
| 33192 | 1148 | (kk_rel_if (kk_all rs1 (kk_rel_eq r2 false_atom)) | 
| 34126 | 1149 | false_atom KK.None) | 
| 33192 | 1150 | end | 
| 1151 | end | |
| 1152 | | Op2 (Or, _, _, u1, u2) => | |
| 1153 | if is_opt_rep (rep_of u1) then kk_rel_if (to_f u2) true_atom (to_r u1) | |
| 1154 | else kk_rel_if (to_f u1) true_atom (to_r u2) | |
| 1155 | | Op2 (And, _, _, u1, u2) => | |
| 1156 | if is_opt_rep (rep_of u1) then kk_rel_if (to_f u2) (to_r u1) false_atom | |
| 1157 | else kk_rel_if (to_f u1) (to_r u2) false_atom | |
| 1158 | | Op2 (Less, _, _, u1, u2) => | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1159 | (case type_of u1 of | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1160 |            @{typ nat} =>
 | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1161 | if is_Cst Unrep u1 then to_compare_with_unrep u2 false_atom | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1162 | else if is_Cst Unrep u2 then to_compare_with_unrep u1 true_atom | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1163 | else kk_nat_less (to_integer u1) (to_integer u2) | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1164 |          | @{typ int} => kk_int_less (to_integer u1) (to_integer u2)
 | 
| 36127 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1165 | | _ => | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1166 | let | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1167 | val R1 = Opt (Atom (card_of_rep (rep_of u1), | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1168 | offset_of_type ofs (type_of u1))) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1169 | in | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1170 | double_rel_rel_let kk | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1171 | (fn r1 => fn r2 => | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1172 | kk_rel_if | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1173 | (fold kk_and (map_filter (fn (u, r) => | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1174 | if is_opt_rep (rep_of u) then SOME (kk_some r) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1175 | else NONE) [(u1, r1), (u2, r2)]) KK.True) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1176 | (atom_from_formula kk bool_j0 (KK.LT (pairself | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1177 | (int_expr_from_atom kk (type_of u1)) (r1, r2)))) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1178 | KK.None) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1179 | (to_rep R1 u1) (to_rep R1 u2) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1180 | end) | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1181 | | Op2 (Triad, _, Opt (Atom (2, j0)), u1, u2) => | 
| 33192 | 1182 | let | 
| 1183 | val f1 = to_f u1 | |
| 1184 | val f2 = to_f u2 | |
| 1185 | in | |
| 1186 | if f1 = f2 then | |
| 1187 | atom_from_formula kk j0 f1 | |
| 1188 | else | |
| 34126 | 1189 | kk_union (kk_rel_if f1 true_atom KK.None) | 
| 1190 | (kk_rel_if f2 KK.None false_atom) | |
| 33192 | 1191 | end | 
| 1192 | | Op2 (Composition, _, R, u1, u2) => | |
| 1193 | let | |
| 33863 
fb13147a3050
generate arguments of relational composition in the right order in Nitpick
 blanchet parents: 
33854diff
changeset | 1194 | val (a_T, b_T) = HOLogic.dest_prodT (domain_type (type_of u1)) | 
| 
fb13147a3050
generate arguments of relational composition in the right order in Nitpick
 blanchet parents: 
33854diff
changeset | 1195 | val (_, c_T) = HOLogic.dest_prodT (domain_type (type_of u2)) | 
| 
fb13147a3050
generate arguments of relational composition in the right order in Nitpick
 blanchet parents: 
33854diff
changeset | 1196 | val ab_k = card_of_domain_from_rep 2 (rep_of u1) | 
| 
fb13147a3050
generate arguments of relational composition in the right order in Nitpick
 blanchet parents: 
33854diff
changeset | 1197 | val bc_k = card_of_domain_from_rep 2 (rep_of u2) | 
| 33192 | 1198 | val ac_k = card_of_domain_from_rep 2 R | 
| 1199 | val a_k = exact_root 2 (ac_k * ab_k div bc_k) | |
| 1200 | val b_k = exact_root 2 (ab_k * bc_k div ac_k) | |
| 1201 | val c_k = exact_root 2 (bc_k * ac_k div ab_k) | |
| 1202 | val a_R = Atom (a_k, offset_of_type ofs a_T) | |
| 1203 | val b_R = Atom (b_k, offset_of_type ofs b_T) | |
| 1204 | val c_R = Atom (c_k, offset_of_type ofs c_T) | |
| 1205 | val body_R = body_rep R | |
| 1206 | in | |
| 1207 | (case body_R of | |
| 1208 | Formula Neut => | |
| 33863 
fb13147a3050
generate arguments of relational composition in the right order in Nitpick
 blanchet parents: 
33854diff
changeset | 1209 | kk_join (to_rep (Func (Struct [a_R, b_R], Formula Neut)) u1) | 
| 
fb13147a3050
generate arguments of relational composition in the right order in Nitpick
 blanchet parents: 
33854diff
changeset | 1210 | (to_rep (Func (Struct [b_R, c_R], Formula Neut)) u2) | 
| 33192 | 1211 | | Opt (Atom (2, _)) => | 
| 1212 | let | |
| 33886 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1213 | (* FIXME: merge with similar code above *) | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1214 | fun must R1 R2 u = | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1215 | kk_join (to_rep (Func (Struct [R1, R2], body_R)) u) true_atom | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1216 | fun may R1 R2 u = | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1217 | kk_difference | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1218 | (full_rel_for_rep (Struct [R1, R2])) | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1219 | (kk_join (to_rep (Func (Struct [R1, R2], body_R)) u) | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1220 | false_atom) | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1221 | in | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1222 | kk_union | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1223 | (kk_product (kk_join (must a_R b_R u1) (must b_R c_R u2)) | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1224 | true_atom) | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1225 | (kk_product (kk_difference | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1226 | (full_rel_for_rep (Struct [a_R, c_R])) | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1227 | (kk_join (may a_R b_R u1) (may b_R c_R u2))) | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1228 | false_atom) | 
| 
cde73f8dbe4e
fix soundness bug in Nitpick's Kodkod generator for the relational composition case
 blanchet parents: 
33863diff
changeset | 1229 | end | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1230 |            | _ => raise NUT ("Nitpick_Kodkod.to_r (Composition)", [u]))
 | 
| 33192 | 1231 | |> rel_expr_from_rel_expr kk R (Func (Struct [a_R, c_R], body_R)) | 
| 1232 | end | |
| 1233 |       | Op2 (Apply, @{typ nat}, _,
 | |
| 1234 | Op2 (Apply, _, _, Cst (Subtract, _, _), u1), u2) => | |
| 1235 | if is_Cst Unrep u2 andalso not (is_opt_rep (rep_of u1)) then | |
| 34126 | 1236 | KK.Atom (offset_of_type ofs nat_T) | 
| 33192 | 1237 | else | 
| 34126 | 1238 | fold kk_join (map to_integer [u1, u2]) (KK.Rel nat_subtract_rel) | 
| 35312 | 1239 | | Op2 (Apply, _, R, u1, u2) => to_apply R u1 u2 | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1240 | | Op2 (Lambda, _, R as Opt (Atom (1, j0)), u1, u2) => | 
| 34126 | 1241 | to_guard [u1, u2] R (KK.Atom j0) | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1242 | | Op2 (Lambda, _, Func (_, Formula Neut), u1, u2) => | 
| 33192 | 1243 | kk_comprehension (untuple to_decl u1) (to_f u2) | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1244 | | Op2 (Lambda, _, Func (_, R2), u1, u2) => | 
| 33192 | 1245 | let | 
| 1246 | val dom_decls = untuple to_decl u1 | |
| 1247 | val ran_schema = atom_schema_of_rep R2 | |
| 1248 | val ran_decls = decls_for_atom_schema ~1 ran_schema | |
| 1249 | val ran_vars = unary_var_seq ~1 (length ran_decls) | |
| 1250 | in | |
| 1251 | kk_comprehension (dom_decls @ ran_decls) | |
| 1252 | (kk_subset (fold1 kk_product ran_vars) | |
| 1253 | (to_rep R2 u2)) | |
| 1254 | end | |
| 1255 | | Op3 (Let, _, R, u1, u2, u3) => | |
| 1256 | kk_rel_let [to_expr_assign u1 u2] (to_rep R u3) | |
| 1257 | | Op3 (If, _, R, u1, u2, u3) => | |
| 1258 | if is_opt_rep (rep_of u1) then | |
| 35386 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 blanchet parents: 
35385diff
changeset | 1259 | triple_rel_rel_let kk | 
| 33192 | 1260 | (fn r1 => fn r2 => fn r3 => | 
| 1261 | let val empty_r = empty_rel_for_rep R in | |
| 1262 | fold1 kk_union | |
| 1263 | [kk_rel_if (kk_rel_eq r1 true_atom) r2 empty_r, | |
| 1264 | kk_rel_if (kk_rel_eq r1 false_atom) r3 empty_r, | |
| 1265 | kk_rel_if (kk_rel_eq r2 r3) | |
| 1266 | (if inline_rel_expr r2 then r2 else r3) empty_r] | |
| 1267 | end) | |
| 1268 | (to_r u1) (to_rep R u2) (to_rep R u3) | |
| 1269 | else | |
| 1270 | kk_rel_if (to_f u1) (to_rep R u2) (to_rep R u3) | |
| 1271 | | Tuple (_, R, us) => | |
| 1272 | (case unopt_rep R of | |
| 1273 | Struct Rs => to_product Rs us | |
| 1274 | | Vect (k, R) => to_product (replicate k R) us | |
| 1275 | | Atom (1, j0) => | |
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 1276 | kk_rel_if (kk_some (fold1 kk_product (map to_r us))) | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 1277 | (KK.Atom j0) KK.None | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1278 |          | _ => raise NUT ("Nitpick_Kodkod.to_r (Tuple)", [u]))
 | 
| 33192 | 1279 | | Construct ([u'], _, _, []) => to_r u' | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1280 | | Construct (discr_u :: sel_us, _, _, arg_us) => | 
| 33192 | 1281 | let | 
| 1282 | val set_rs = | |
| 1283 | map2 (fn sel_u => fn arg_u => | |
| 1284 | let | |
| 1285 | val (R1, R2) = dest_Func (rep_of sel_u) | |
| 1286 | val sel_r = to_r sel_u | |
| 1287 | val arg_r = to_opt R2 arg_u | |
| 1288 | in | |
| 1289 | if is_one_rep R2 then | |
| 1290 | kk_n_fold_join kk true R2 R1 arg_r | |
| 1291 | (kk_project sel_r (flip_nums (arity_of_rep R2))) | |
| 1292 | else | |
| 34288 
cf455b5880e1
reduced arity of Nitpick selectors associated with sets by 1, by using "Formula" instead of "Atom 2"
 blanchet parents: 
34126diff
changeset | 1293 | kk_comprehension [KK.DeclOne ((1, ~1), to_r discr_u)] | 
| 34126 | 1294 | (kk_rel_eq (kk_join (KK.Var (1, ~1)) sel_r) arg_r) | 
| 35695 | 1295 | |> is_opt_rep (rep_of arg_u) ? to_guard [arg_u] R1 | 
| 33192 | 1296 | end) sel_us arg_us | 
| 1297 | in fold1 kk_intersect set_rs end | |
| 34126 | 1298 | | BoundRel (x, _, _, _) => KK.Var x | 
| 1299 | | FreeRel (x, _, _, _) => KK.Rel x | |
| 1300 | | RelReg (j, _, R) => KK.RelReg (arity_of_rep R, j) | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1301 |       | u => raise NUT ("Nitpick_Kodkod.to_r", [u])
 | 
| 33192 | 1302 | and to_decl (BoundRel (x, _, R, _)) = | 
| 34126 | 1303 | KK.DeclOne (x, KK.AtomSeq (the_single (atom_schema_of_rep R))) | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1304 |       | to_decl u = raise NUT ("Nitpick_Kodkod.to_decl", [u])
 | 
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1305 | and to_expr_assign (FormulaReg (j, _, _)) u = | 
| 34126 | 1306 | KK.AssignFormulaReg (j, to_f u) | 
| 33192 | 1307 | | to_expr_assign (RelReg (j, _, R)) u = | 
| 34126 | 1308 | KK.AssignRelReg ((arity_of_rep R, j), to_r u) | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1309 |       | to_expr_assign u1 _ = raise NUT ("Nitpick_Kodkod.to_expr_assign", [u1])
 | 
| 39345 | 1310 | and to_atom (x as (_, j0)) u = | 
| 33192 | 1311 | case rep_of u of | 
| 1312 | Formula _ => atom_from_formula kk j0 (to_f u) | |
| 1313 | | R => atom_from_rel_expr kk x R (to_r u) | |
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 1314 | and to_struct Rs u = struct_from_rel_expr kk Rs (rep_of u) (to_r u) | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 1315 | and to_vect k R u = vect_from_rel_expr kk k R (rep_of u) (to_r u) | 
| 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 1316 | and to_func R1 R2 u = rel_expr_to_func kk R1 R2 (rep_of u) (to_r u) | 
| 33192 | 1317 | and to_opt R u = | 
| 1318 | let val old_R = rep_of u in | |
| 1319 | if is_opt_rep old_R then | |
| 1320 | rel_expr_from_rel_expr kk (Opt R) old_R (to_r u) | |
| 1321 | else | |
| 1322 | to_rep R u | |
| 1323 | end | |
| 1324 | and to_rep (Atom x) u = to_atom x u | |
| 1325 | | to_rep (Struct Rs) u = to_struct Rs u | |
| 1326 | | to_rep (Vect (k, R)) u = to_vect k R u | |
| 1327 | | to_rep (Func (R1, R2)) u = to_func R1 R2 u | |
| 1328 | | to_rep (Opt R) u = to_opt R u | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1329 |       | to_rep R _ = raise REP ("Nitpick_Kodkod.to_rep", [R])
 | 
| 33192 | 1330 | and to_integer u = to_opt (one_rep ofs (type_of u) (rep_of u)) u | 
| 1331 | and to_guard guard_us R r = | |
| 1332 | let | |
| 1333 | val unpacked_rs = unpack_joins r | |
| 1334 | val plain_guard_rs = | |
| 1335 | map to_r (filter (is_Opt o rep_of) guard_us) | |
| 1336 | |> filter_out (member (op =) unpacked_rs) | |
| 1337 | val func_guard_us = | |
| 1338 | filter ((is_Func andf is_opt_rep) o rep_of) guard_us | |
| 1339 | val func_guard_rs = map to_r func_guard_us | |
| 1340 | val guard_fs = | |
| 1341 | map kk_no plain_guard_rs @ | |
| 1342 | map2 (kk_not oo kk_n_ary_function kk) | |
| 1343 | (map (unopt_rep o rep_of) func_guard_us) func_guard_rs | |
| 1344 | in | |
| 35695 | 1345 | if null guard_fs then r | 
| 1346 | else kk_rel_if (fold1 kk_or guard_fs) (empty_rel_for_rep R) r | |
| 33192 | 1347 | end | 
| 1348 | and to_project new_R old_R r j0 = | |
| 1349 | rel_expr_from_rel_expr kk new_R old_R | |
| 1350 | (kk_project_seq r j0 (arity_of_rep old_R)) | |
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 1351 | and to_product Rs us = fold1 kk_product (map (uncurry to_opt) (Rs ~~ us)) | 
| 39345 | 1352 | and to_nth_pair_sel n res_R u = | 
| 33192 | 1353 | case u of | 
| 1354 | Tuple (_, _, us) => to_rep res_R (nth us n) | |
| 1355 | | _ => let | |
| 1356 | val R = rep_of u | |
| 1357 | val (a_T, b_T) = HOLogic.dest_prodT (type_of u) | |
| 1358 | val Rs = | |
| 1359 | case unopt_rep R of | |
| 1360 | Struct (Rs as [_, _]) => Rs | |
| 1361 | | _ => | |
| 1362 | let | |
| 1363 | val res_card = card_of_rep res_R | |
| 1364 | val other_card = card_of_rep R div res_card | |
| 1365 | val (a_card, b_card) = (res_card, other_card) | |
| 1366 | |> n = 1 ? swap | |
| 1367 | in | |
| 1368 | [Atom (a_card, offset_of_type ofs a_T), | |
| 1369 | Atom (b_card, offset_of_type ofs b_T)] | |
| 1370 | end | |
| 1371 | val nth_R = nth Rs n | |
| 1372 | val j0 = if n = 0 then 0 else arity_of_rep (hd Rs) | |
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 1373 | in to_project res_R nth_R (to_rep (Opt (Struct Rs)) u) j0 end | 
| 33192 | 1374 | and to_set_bool_op connective set_oper u1 u2 = | 
| 1375 | let | |
| 1376 | val min_R = min_rep (rep_of u1) (rep_of u2) | |
| 1377 | val r1 = to_rep min_R u1 | |
| 1378 | val r2 = to_rep min_R u2 | |
| 1379 | in | |
| 1380 | case min_R of | |
| 1381 | Vect (k, Atom _) => kk_vect_set_bool_op connective k r1 r2 | |
| 35280 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1382 | | Func (_, Formula Neut) => set_oper r1 r2 | 
| 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1383 | | Func (_, Atom _) => set_oper (kk_join r1 true_atom) | 
| 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 blanchet parents: 
35220diff
changeset | 1384 | (kk_join r2 true_atom) | 
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 1385 |         | _ => raise REP ("Nitpick_Kodkod.to_set_bool_op", [min_R])
 | 
| 33192 | 1386 | end | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1387 | and to_bit_word_unary_op T R oper = | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1388 | let | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1389 | val Ts = strip_type T ||> single |> op @ | 
| 34126 | 1390 | fun int_arg j = int_expr_from_atom kk (nth Ts j) (KK.Var (1, j)) | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1391 | in | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1392 | kk_comprehension (decls_for_atom_schema 0 (atom_schema_of_rep R)) | 
| 34126 | 1393 | (KK.FormulaLet | 
| 1394 | (map (fn j => KK.AssignIntReg (j, int_arg j)) (0 upto 1), | |
| 1395 | KK.IntEq (KK.IntReg 1, oper (KK.IntReg 0)))) | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1396 | end | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1397 | and to_bit_word_binary_op T R opt_guard opt_oper = | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1398 | let | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1399 | val Ts = strip_type T ||> single |> op @ | 
| 34126 | 1400 | fun int_arg j = int_expr_from_atom kk (nth Ts j) (KK.Var (1, j)) | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1401 | in | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1402 | kk_comprehension (decls_for_atom_schema 0 (atom_schema_of_rep R)) | 
| 34126 | 1403 | (KK.FormulaLet | 
| 1404 | (map (fn j => KK.AssignIntReg (j, int_arg j)) (0 upto 2), | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1405 | fold1 kk_and | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1406 | ((case opt_guard of | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1407 | NONE => [] | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1408 | | SOME guard => | 
| 34126 | 1409 | [guard (KK.IntReg 0) (KK.IntReg 1) (KK.IntReg 2)]) @ | 
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1410 | (case opt_oper of | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1411 | NONE => [] | 
| 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1412 | | SOME oper => | 
| 34126 | 1413 | [KK.IntEq (KK.IntReg 2, | 
| 1414 | oper (KK.IntReg 0) (KK.IntReg 1))])))) | |
| 34124 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 blanchet parents: 
34121diff
changeset | 1415 | end | 
| 39345 | 1416 | and to_apply (R as Formula _) _ _ = | 
| 36127 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1417 |         raise REP ("Nitpick_Kodkod.to_apply", [R])
 | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1418 | | to_apply res_R func_u arg_u = | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1419 | case unopt_rep (rep_of func_u) of | 
| 38190 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 blanchet parents: 
38182diff
changeset | 1420 | Atom (1, j0) => | 
| 33192 | 1421 | to_guard [arg_u] res_R | 
| 36127 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1422 | (rel_expr_from_rel_expr kk res_R (Atom (1, j0)) (to_r func_u)) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1423 | | Atom (k, _) => | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1424 | let | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1425 | val dom_card = card_of_rep (rep_of arg_u) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1426 | val ran_R = Atom (exact_root dom_card k, | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1427 | offset_of_type ofs (range_type (type_of func_u))) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1428 | in | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1429 | to_apply_vect dom_card ran_R res_R (to_vect dom_card ran_R func_u) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1430 | arg_u | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1431 | end | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1432 | | Vect (1, R') => | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1433 | to_guard [arg_u] res_R | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1434 | (rel_expr_from_rel_expr kk res_R R' (to_r func_u)) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1435 | | Vect (k, R') => to_apply_vect k R' res_R (to_r func_u) arg_u | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1436 | | Func (R, Formula Neut) => | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1437 | to_guard [arg_u] res_R (rel_expr_from_formula kk res_R | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1438 | (kk_subset (to_opt R arg_u) (to_r func_u))) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1439 | | Func (R1, R2) => | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1440 | rel_expr_from_rel_expr kk res_R R2 | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1441 | (kk_n_fold_join kk true R1 R2 (to_opt R1 arg_u) (to_r func_u)) | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1442 | |> body_rep R2 = Formula Neut ? to_guard [arg_u] res_R | 
| 
e91292c520be
fix bug in Nitpick's handling of "<" (exposed by "GCD.setprod_coprime_int")
 blanchet parents: 
35695diff
changeset | 1443 |         | _ => raise NUT ("Nitpick_Kodkod.to_apply", [func_u])
 | 
| 33192 | 1444 | and to_apply_vect k R' res_R func_r arg_u = | 
| 1445 | let | |
| 1446 | val arg_R = one_rep ofs (type_of arg_u) (unopt_rep (rep_of arg_u)) | |
| 1447 | val vect_r = vect_from_rel_expr kk k res_R (Vect (k, R')) func_r | |
| 1448 | val vect_rs = unpack_vect_in_chunks kk (arity_of_rep res_R) k vect_r | |
| 1449 | in | |
| 1450 | kk_case_switch kk arg_R res_R (to_opt arg_R arg_u) | |
| 1451 | (all_singletons_for_rep arg_R) vect_rs | |
| 1452 | end | |
| 1453 | and to_could_be_unrep neg u = | |
| 34126 | 1454 | if neg andalso is_opt_rep (rep_of u) then kk_no (to_r u) else KK.False | 
| 33192 | 1455 | and to_compare_with_unrep u r = | 
| 33892 
3937da7e13ea
fixed arity of some empty relations in Nitpick's Kodkod generator;
 blanchet parents: 
33886diff
changeset | 1456 | if is_opt_rep (rep_of u) then | 
| 
3937da7e13ea
fixed arity of some empty relations in Nitpick's Kodkod generator;
 blanchet parents: 
33886diff
changeset | 1457 | kk_rel_if (kk_some (to_r u)) r (empty_rel_for_rep (rep_of u)) | 
| 
3937da7e13ea
fixed arity of some empty relations in Nitpick's Kodkod generator;
 blanchet parents: 
33886diff
changeset | 1458 | else | 
| 
3937da7e13ea
fixed arity of some empty relations in Nitpick's Kodkod generator;
 blanchet parents: 
33886diff
changeset | 1459 | r | 
| 33192 | 1460 | in to_f_with_polarity Pos u end | 
| 1461 | ||
| 41802 | 1462 | fun declarative_axiom_for_plain_rel kk (FreeRel (x, _, R as Func _, nick)) = | 
| 1463 |     kk_n_ary_function kk (R |> nick = @{const_name List.set} ? unopt_rep)
 | |
| 1464 | (KK.Rel x) | |
| 1465 |   | declarative_axiom_for_plain_rel ({kk_lone, kk_one, ...} : kodkod_constrs)
 | |
| 1466 | (FreeRel (x, _, R, _)) = | |
| 1467 | if is_one_rep R then kk_one (KK.Rel x) | |
| 1468 | else if is_lone_rep R andalso card_of_rep R > 1 then kk_lone (KK.Rel x) | |
| 1469 | else KK.True | |
| 1470 | | declarative_axiom_for_plain_rel _ u = | |
| 1471 |     raise NUT ("Nitpick_Kodkod.declarative_axiom_for_plain_rel", [u])
 | |
| 1472 | ||
| 1473 | fun const_triple rel_table (x as (s, T)) = | |
| 1474 | case the_name rel_table (ConstName (s, T, Any)) of | |
| 1475 | FreeRel ((n, j), _, R, _) => (KK.Rel (n, j), R, n) | |
| 1476 |   | _ => raise TERM ("Nitpick_Kodkod.const_triple", [Const x])
 | |
| 1477 | ||
| 1478 | fun discr_rel_expr rel_table = #1 o const_triple rel_table o discr_for_constr | |
| 1479 | ||
| 1480 | fun nfa_transitions_for_sel hol_ctxt binarize | |
| 1481 |                             ({kk_project, ...} : kodkod_constrs) rel_table
 | |
| 1482 | (dtypes : datatype_spec list) constr_x n = | |
| 1483 | let | |
| 1484 | val x as (_, T) = | |
| 1485 | binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize constr_x n | |
| 1486 | val (r, R, arity) = const_triple rel_table x | |
| 1487 | val type_schema = type_schema_of_rep T R | |
| 1488 | in | |
| 1489 | map_filter (fn (j, T) => | |
| 1490 | if forall (not_equal T o #typ) dtypes then NONE | |
| 1491 | else SOME ((x, kk_project r (map KK.Num [0, j])), T)) | |
| 1492 | (index_seq 1 (arity - 1) ~~ tl type_schema) | |
| 1493 | end | |
| 1494 | fun nfa_transitions_for_constr hol_ctxt binarize kk rel_table dtypes | |
| 1495 | (x as (_, T)) = | |
| 1496 | maps (nfa_transitions_for_sel hol_ctxt binarize kk rel_table dtypes x) | |
| 1497 | (index_seq 0 (num_sels_for_constr_type T)) | |
| 1498 | fun nfa_entry_for_datatype _ _ _ _ _ ({co = true, ...} : datatype_spec) = NONE
 | |
| 1499 |   | nfa_entry_for_datatype _ _ _ _ _ {standard = false, ...} = NONE
 | |
| 1500 |   | nfa_entry_for_datatype _ _ _ _ _ {deep = false, ...} = NONE
 | |
| 1501 | | nfa_entry_for_datatype hol_ctxt binarize kk rel_table dtypes | |
| 1502 |                            {typ, constrs, ...} =
 | |
| 1503 | SOME (typ, maps (nfa_transitions_for_constr hol_ctxt binarize kk rel_table | |
| 1504 | dtypes o #const) constrs) | |
| 1505 | ||
| 1506 | val empty_rel = KK.Product (KK.None, KK.None) | |
| 1507 | ||
| 1508 | fun direct_path_rel_exprs nfa start_T final_T = | |
| 1509 | case AList.lookup (op =) nfa final_T of | |
| 1510 | SOME trans => map (snd o fst) (filter (curry (op =) start_T o snd) trans) | |
| 1511 | | NONE => [] | |
| 1512 | and any_path_rel_expr ({kk_union, ...} : kodkod_constrs) nfa [] start_T
 | |
| 1513 | final_T = | |
| 1514 | fold kk_union (direct_path_rel_exprs nfa start_T final_T) | |
| 1515 | (if start_T = final_T then KK.Iden else empty_rel) | |
| 1516 |   | any_path_rel_expr (kk as {kk_union, ...}) nfa (T :: Ts) start_T final_T =
 | |
| 1517 | kk_union (any_path_rel_expr kk nfa Ts start_T final_T) | |
| 1518 | (knot_path_rel_expr kk nfa Ts start_T T final_T) | |
| 1519 | and knot_path_rel_expr (kk as {kk_join, kk_reflexive_closure, ...}) nfa Ts
 | |
| 1520 | start_T knot_T final_T = | |
| 1521 | kk_join (kk_join (any_path_rel_expr kk nfa Ts knot_T final_T) | |
| 1522 | (kk_reflexive_closure (loop_path_rel_expr kk nfa Ts knot_T))) | |
| 1523 | (any_path_rel_expr kk nfa Ts start_T knot_T) | |
| 1524 | and loop_path_rel_expr ({kk_union, ...} : kodkod_constrs) nfa [] start_T =
 | |
| 1525 | fold kk_union (direct_path_rel_exprs nfa start_T start_T) empty_rel | |
| 1526 |   | loop_path_rel_expr (kk as {kk_union, kk_closure, ...}) nfa (T :: Ts)
 | |
| 1527 | start_T = | |
| 1528 | if start_T = T then | |
| 1529 | kk_closure (loop_path_rel_expr kk nfa Ts start_T) | |
| 1530 | else | |
| 1531 | kk_union (loop_path_rel_expr kk nfa Ts start_T) | |
| 1532 | (knot_path_rel_expr kk nfa Ts start_T T start_T) | |
| 1533 | ||
| 1534 | fun add_nfa_to_graph [] = I | |
| 1535 | | add_nfa_to_graph ((_, []) :: nfa) = add_nfa_to_graph nfa | |
| 1536 | | add_nfa_to_graph ((T, ((_, T') :: transitions)) :: nfa) = | |
| 1537 | add_nfa_to_graph ((T, transitions) :: nfa) o Typ_Graph.add_edge (T, T') o | |
| 1538 | Typ_Graph.default_node (T', ()) o Typ_Graph.default_node (T, ()) | |
| 1539 | ||
| 1540 | fun strongly_connected_sub_nfas nfa = | |
| 1541 | add_nfa_to_graph nfa Typ_Graph.empty | |
| 1542 | |> Typ_Graph.strong_conn | |
| 1543 | |> map (fn keys => filter (member (op =) keys o fst) nfa) | |
| 1544 | ||
| 1545 | fun acyclicity_axiom_for_datatype (kk as {kk_no, kk_intersect, ...}) nfa
 | |
| 1546 | start_T = | |
| 1547 | kk_no (kk_intersect | |
| 1548 | (loop_path_rel_expr kk nfa (pull start_T (map fst nfa)) start_T) | |
| 1549 | KK.Iden) | |
| 1550 | (* Cycle breaking in the bounds takes care of singly recursive datatypes, hence | |
| 1551 | the first equation. *) | |
| 1552 | fun acyclicity_axioms_for_datatypes _ [_] = [] | |
| 1553 | | acyclicity_axioms_for_datatypes kk nfas = | |
| 1554 | maps (fn nfa => map (acyclicity_axiom_for_datatype kk nfa o fst) nfa) nfas | |
| 1555 | ||
| 1556 | fun atom_equation_for_nut ofs kk (u, j) = | |
| 1557 | let val dummy_u = RelReg (0, type_of u, rep_of u) in | |
| 1558 | case Op2 (DefEq, bool_T, Formula Pos, dummy_u, u) | |
| 1559 | |> kodkod_formula_from_nut ofs kk of | |
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1560 | KK.RelEq (KK.RelReg _, r) => SOME (KK.RelEq (KK.Atom j, r)) | 
| 41802 | 1561 |     | _ => raise BAD ("Nitpick_Kodkod.atom_equation_for_nut",
 | 
| 1562 | "malformed Kodkod formula") | |
| 1563 | end | |
| 1564 | ||
| 41995 | 1565 | fun needed_values_for_datatype [] _ _ = SOME [] | 
| 1566 | | needed_values_for_datatype need_us ofs | |
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1567 |                                ({typ, card, constrs, ...} : datatype_spec) =
 | 
| 41802 | 1568 | let | 
| 1569 | fun aux (u as Construct (FreeRel (_, _, _, s) :: _, T, _, us)) = | |
| 1570 | fold aux us | |
| 1571 | #> (fn NONE => NONE | |
| 1572 | | accum as SOME (loose, fixed) => | |
| 1573 | if T = typ then | |
| 1574 | case AList.lookup (op =) fixed u of | |
| 1575 | SOME _ => accum | |
| 1576 | | NONE => | |
| 1577 | let | |
| 1578 | val constr_s = constr_name_for_sel_like s | |
| 1579 |                        val {delta, epsilon, ...} =
 | |
| 1580 | constrs | |
| 1581 |                          |> List.find (fn {const, ...} => fst const = constr_s)
 | |
| 1582 | |> the | |
| 1583 | val j0 = offset_of_type ofs T | |
| 1584 | in | |
| 1585 | case find_first (fn j => j >= delta andalso | |
| 1586 | j < delta + epsilon) loose of | |
| 1587 | SOME j => | |
| 1588 | SOME (remove (op =) j loose, (u, j0 + j) :: fixed) | |
| 1589 | | NONE => NONE | |
| 1590 | end | |
| 1591 | else | |
| 1592 | accum) | |
| 41985 | 1593 | | aux _ = I | 
| 41996 | 1594 | in | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1595 | SOME (index_seq 0 card, []) |> fold aux need_us |> Option.map (rev o snd) | 
| 41996 | 1596 | end | 
| 41995 | 1597 | |
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1598 | fun needed_value_axioms_for_datatype _ _ _ (_, NONE) = [KK.False] | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1599 | | needed_value_axioms_for_datatype ofs kk dtypes (T, SOME fixed) = | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1600 | if is_datatype_nat_like (the (datatype_spec dtypes T)) then [] | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1601 | else fixed |> map_filter (atom_equation_for_nut ofs kk) | 
| 41802 | 1602 | |
| 1603 | fun all_ge ({kk_join, kk_reflexive_closure, ...} : kodkod_constrs) z r =
 | |
| 1604 | kk_join r (kk_reflexive_closure (KK.Rel (suc_rel_for_atom_seq z))) | |
| 1605 | fun gt ({kk_subset, kk_join, kk_closure, ...} : kodkod_constrs) z r1 r2 =
 | |
| 1606 | kk_subset r1 (kk_join r2 (kk_closure (KK.Rel (suc_rel_for_atom_seq z)))) | |
| 1607 | ||
| 1608 | fun constr_quadruple ({const = (s, T), delta, epsilon, ...} : constr_spec) =
 | |
| 1609 | (delta, (epsilon, (num_binder_types T, s))) | |
| 1610 | val constr_ord = | |
| 1611 | prod_ord int_ord (prod_ord int_ord (prod_ord int_ord string_ord)) | |
| 1612 | o pairself constr_quadruple | |
| 1613 | ||
| 1614 | fun datatype_ord (({card = card1, self_rec = self_rec1, constrs = constr1, ...},
 | |
| 1615 |                    {card = card2, self_rec = self_rec2, constrs = constr2, ...})
 | |
| 1616 | : datatype_spec * datatype_spec) = | |
| 1617 | prod_ord int_ord (prod_ord bool_ord int_ord) | |
| 1618 | ((card1, (self_rec1, length constr1)), | |
| 1619 | (card2, (self_rec2, length constr2))) | |
| 1620 | ||
| 1621 | (* We must absolutely tabulate "suc" for all datatypes whose selector bounds | |
| 1622 | break cycles; otherwise, we may end up with two incompatible symmetry | |
| 1623 | breaking orders, leading to spurious models. *) | |
| 1624 | fun should_tabulate_suc_for_type dtypes T = | |
| 1625 | is_asymmetric_nondatatype T orelse | |
| 1626 | case datatype_spec dtypes T of | |
| 1627 |     SOME {self_rec, ...} => self_rec
 | |
| 1628 | | NONE => false | |
| 1629 | ||
| 1630 | fun lex_order_rel_expr (kk as {kk_implies, kk_and, kk_subset, kk_join, ...})
 | |
| 1631 | dtypes sel_quadruples = | |
| 1632 | case sel_quadruples of | |
| 1633 | [] => KK.True | |
| 1634 | | ((r, Func (Atom _, Atom x), 2), (_, Type (_, [_, T]))) :: sel_quadruples' => | |
| 1635 | let val z = (x, should_tabulate_suc_for_type dtypes T) in | |
| 1636 | if null sel_quadruples' then | |
| 1637 | gt kk z (kk_join (KK.Var (1, 1)) r) (kk_join (KK.Var (1, 0)) r) | |
| 1638 | else | |
| 1639 | kk_and (kk_subset (kk_join (KK.Var (1, 1)) r) | |
| 1640 | (all_ge kk z (kk_join (KK.Var (1, 0)) r))) | |
| 1641 | (kk_implies (kk_subset (kk_join (KK.Var (1, 1)) r) | |
| 1642 | (kk_join (KK.Var (1, 0)) r)) | |
| 1643 | (lex_order_rel_expr kk dtypes sel_quadruples')) | |
| 1644 | end | |
| 1645 | (* Skip constructors components that aren't atoms, since we cannot compare | |
| 1646 | these easily. *) | |
| 1647 | | _ :: sel_quadruples' => lex_order_rel_expr kk dtypes sel_quadruples' | |
| 1648 | ||
| 1649 | fun is_nil_like_constr_type dtypes T = | |
| 1650 | case datatype_spec dtypes T of | |
| 1651 |     SOME {constrs, ...} =>
 | |
| 1652 | (case filter_out (is_self_recursive_constr_type o snd o #const) constrs of | |
| 1653 |        [{const = (_, T'), ...}] => T = T'
 | |
| 1654 | | _ => false) | |
| 1655 | | NONE => false | |
| 1656 | ||
| 1657 | fun sym_break_axioms_for_constr_pair hol_ctxt binarize | |
| 1658 |        (kk as {kk_all, kk_or, kk_implies, kk_and, kk_some, kk_intersect,
 | |
| 1659 | kk_join, ...}) rel_table nfas dtypes | |
| 1660 | (constr_ord, | |
| 1661 |         ({const = const1 as (_, T1), delta = delta1, epsilon = epsilon1, ...},
 | |
| 1662 |          {const = const2 as (_, _), delta = delta2, epsilon = epsilon2, ...})
 | |
| 1663 | : constr_spec * constr_spec) = | |
| 1664 | let | |
| 1665 | val dataT = body_type T1 | |
| 1666 | val nfa = nfas |> find_first (exists (curry (op =) dataT o fst)) |> these | |
| 1667 | val rec_Ts = nfa |> map fst | |
| 1668 | fun rec_and_nonrec_sels (x as (_, T)) = | |
| 1669 | index_seq 0 (num_sels_for_constr_type T) | |
| 1670 | |> map (binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize x) | |
| 1671 | |> List.partition (member (op =) rec_Ts o range_type o snd) | |
| 1672 | val sel_xs1 = rec_and_nonrec_sels const1 |> op @ | |
| 1673 | in | |
| 1674 | if constr_ord = EQUAL andalso null sel_xs1 then | |
| 1675 | [] | |
| 1676 | else | |
| 1677 | let | |
| 1678 | val z = | |
| 1679 | (case #2 (const_triple rel_table (discr_for_constr const1)) of | |
| 1680 | Func (Atom x, Formula _) => x | |
| 1681 |            | R => raise REP ("Nitpick_Kodkod.sym_break_axioms_for_constr_pair",
 | |
| 1682 | [R]), should_tabulate_suc_for_type dtypes dataT) | |
| 1683 | val (rec_sel_xs2, nonrec_sel_xs2) = rec_and_nonrec_sels const2 | |
| 1684 | val sel_xs2 = rec_sel_xs2 @ nonrec_sel_xs2 | |
| 1685 | fun sel_quadruples2 () = sel_xs2 |> map (`(const_triple rel_table)) | |
| 1686 | (* If the two constructors are the same, we drop the first selector | |
| 1687 | because that one is always checked by the lexicographic order. | |
| 1688 | We sometimes also filter out direct subterms, because those are | |
| 1689 | already handled by the acyclicity breaking in the bound | |
| 1690 | declarations. *) | |
| 1691 | fun filter_out_sels no_direct sel_xs = | |
| 1692 | apsnd (filter_out | |
| 1693 | (fn ((x, _), T) => | |
| 1694 | (constr_ord = EQUAL andalso x = hd sel_xs) orelse | |
| 1695 | (T = dataT andalso | |
| 1696 | (no_direct orelse not (member (op =) sel_xs x))))) | |
| 1697 | fun subterms_r no_direct sel_xs j = | |
| 1698 | loop_path_rel_expr kk (map (filter_out_sels no_direct sel_xs) nfa) | |
| 1699 | (filter_out (curry (op =) dataT) (map fst nfa)) dataT | |
| 1700 | |> kk_join (KK.Var (1, j)) | |
| 1701 | in | |
| 1702 | [kk_all [KK.DeclOne ((1, 0), discr_rel_expr rel_table const1), | |
| 1703 | KK.DeclOne ((1, 1), discr_rel_expr rel_table const2)] | |
| 1704 | (kk_implies | |
| 1705 | (if delta2 >= epsilon1 then KK.True | |
| 1706 | else if delta1 >= epsilon2 - 1 then KK.False | |
| 1707 | else gt kk z (KK.Var (1, 1)) (KK.Var (1, 0))) | |
| 1708 | (kk_or | |
| 1709 | (if is_nil_like_constr_type dtypes T1 then | |
| 1710 | KK.True | |
| 1711 | else | |
| 1712 | kk_some (kk_intersect (subterms_r false sel_xs2 1) | |
| 1713 | (all_ge kk z (KK.Var (1, 0))))) | |
| 1714 | (case constr_ord of | |
| 1715 | EQUAL => | |
| 1716 | kk_and | |
| 1717 | (lex_order_rel_expr kk dtypes (sel_quadruples2 ())) | |
| 1718 | (kk_all [KK.DeclOne ((1, 2), | |
| 1719 | subterms_r true sel_xs1 0)] | |
| 1720 | (gt kk z (KK.Var (1, 1)) (KK.Var (1, 2)))) | |
| 1721 | | LESS => | |
| 1722 | kk_all [KK.DeclOne ((1, 2), | |
| 1723 | subterms_r false sel_xs1 0)] | |
| 1724 | (gt kk z (KK.Var (1, 1)) (KK.Var (1, 2))) | |
| 1725 | | GREATER => KK.False)))] | |
| 1726 | end | |
| 1727 | end | |
| 1728 | ||
| 1729 | fun sym_break_axioms_for_datatype hol_ctxt binarize kk rel_table nfas dtypes | |
| 1730 |                                   ({constrs, ...} : datatype_spec) =
 | |
| 1731 | let | |
| 1732 | val constrs = sort constr_ord constrs | |
| 1733 | val constr_pairs = all_distinct_unordered_pairs_of constrs | |
| 1734 | in | |
| 1735 | map (pair EQUAL) (constrs ~~ constrs) @ | |
| 1736 | map (pair LESS) constr_pairs @ | |
| 1737 | map (pair GREATER) (map swap constr_pairs) | |
| 1738 | |> maps (sym_break_axioms_for_constr_pair hol_ctxt binarize kk rel_table | |
| 1739 | nfas dtypes) | |
| 1740 | end | |
| 1741 | ||
| 41875 | 1742 | fun is_datatype_in_needed_value T (Construct (_, T', _, us)) = | 
| 1743 | T = T' orelse exists (is_datatype_in_needed_value T) us | |
| 1744 | | is_datatype_in_needed_value _ _ = false | |
| 41802 | 1745 | |
| 1746 | val min_sym_break_card = 7 | |
| 1747 | ||
| 41875 | 1748 | fun sym_break_axioms_for_datatypes hol_ctxt binarize need_us datatype_sym_break | 
| 1749 | kk rel_table nfas dtypes = | |
| 41802 | 1750 | if datatype_sym_break = 0 then | 
| 1751 | [] | |
| 1752 | else | |
| 41803 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1753 | dtypes |> filter is_datatype_acyclic | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1754 |            |> filter (fn {constrs = [_], ...} => false
 | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1755 |                        | {card, constrs, ...} =>
 | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1756 | card >= min_sym_break_card andalso | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1757 | forall (forall (not o is_higher_order_type) | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1758 | o binder_types o snd o #const) constrs) | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1759 | |> filter_out | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1760 |                   (fn {typ, ...} =>
 | 
| 41875 | 1761 | exists (is_datatype_in_needed_value typ) need_us) | 
| 41803 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1762 | |> (fn dtypes' => | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1763 | dtypes' |> length dtypes' > datatype_sym_break | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1764 | ? (sort (datatype_ord o swap) | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1765 | #> take datatype_sym_break)) | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1766 | |> maps (sym_break_axioms_for_datatype hol_ctxt binarize kk rel_table | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1767 | nfas dtypes) | 
| 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 blanchet parents: 
41802diff
changeset | 1768 | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1769 | fun sel_axioms_for_sel hol_ctxt binarize j0 | 
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1770 |         (kk as {kk_all, kk_formula_if, kk_subset, kk_no, kk_join, ...})
 | 
| 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1771 |         need_vals rel_table dom_r (dtype as {typ, ...})
 | 
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1772 |         ({const, delta, epsilon, exclusive, ...} : constr_spec) n =
 | 
| 41802 | 1773 | let | 
| 1774 | val x = binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize const n | |
| 1775 | val (r, R, _) = const_triple rel_table x | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1776 | val rel_x = | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1777 | case r of | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1778 | KK.Rel x => x | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1779 |       | _ => raise BAD ("Nitpick_Kodkod.sel_axioms_for_sel", "non-Rel")
 | 
| 41802 | 1780 | val R2 = dest_Func R |> snd | 
| 1781 | val z = (epsilon - delta, delta + j0) | |
| 1782 | in | |
| 1783 | if exclusive then | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1784 | [kk_n_ary_function kk (Func (Atom z, R2)) r] | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1785 | else if all_values_are_needed need_vals dtype then | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1786 | typ |> needed_values need_vals | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1787 | |> filter (is_sel_of_constr rel_x) | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1788 | |> map (fn (_, j) => kk_n_ary_function kk R2 (kk_join (KK.Atom j) r)) | 
| 41802 | 1789 | else | 
| 1790 | let val r' = kk_join (KK.Var (1, 0)) r in | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1791 | [kk_all [KK.DeclOne ((1, 0), KK.AtomSeq z)] | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1792 | (kk_formula_if (kk_subset (KK.Var (1, 0)) dom_r) | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1793 | (kk_n_ary_function kk R2 r') (kk_no r'))] | 
| 41802 | 1794 | end | 
| 1795 | end | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1796 | fun sel_axioms_for_constr hol_ctxt binarize bits j0 kk need_vals rel_table | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1797 |         dtype (constr as {const, delta, epsilon, explicit_max, ...}) =
 | 
| 41802 | 1798 | let | 
| 1799 | val honors_explicit_max = | |
| 1800 | explicit_max < 0 orelse epsilon - delta <= explicit_max | |
| 1801 | in | |
| 1802 | if explicit_max = 0 then | |
| 1803 | [formula_for_bool honors_explicit_max] | |
| 1804 | else | |
| 1805 | let | |
| 1806 | val dom_r = discr_rel_expr rel_table const | |
| 1807 | val max_axiom = | |
| 1808 | if honors_explicit_max then | |
| 1809 | KK.True | |
| 1810 | else if bits = 0 orelse | |
| 1811 | is_twos_complement_representable bits (epsilon - delta) then | |
| 1812 | KK.LE (KK.Cardinality dom_r, KK.Num explicit_max) | |
| 1813 | else | |
| 1814 |             raise TOO_SMALL ("Nitpick_Kodkod.sel_axioms_for_constr",
 | |
| 1815 | "\"bits\" value " ^ string_of_int bits ^ | |
| 1816 | " too small for \"max\"") | |
| 1817 | in | |
| 1818 | max_axiom :: | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1819 | maps (sel_axioms_for_sel hol_ctxt binarize j0 kk need_vals rel_table | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1820 | dom_r dtype constr) | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1821 | (index_seq 0 (num_sels_for_constr_type (snd const))) | 
| 41802 | 1822 | end | 
| 1823 | end | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1824 | fun sel_axioms_for_datatype hol_ctxt binarize bits j0 kk rel_table need_vals | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1825 |                             (dtype as {constrs, ...}) =
 | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1826 | maps (sel_axioms_for_constr hol_ctxt binarize bits j0 kk rel_table need_vals | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1827 | dtype) constrs | 
| 41802 | 1828 | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1829 | fun uniqueness_axioms_for_constr hol_ctxt binarize | 
| 41802 | 1830 |         ({kk_all, kk_implies, kk_and, kk_rel_eq, kk_lone, kk_join, ...}
 | 
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1831 | : kodkod_constrs) need_vals rel_table dtype | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1832 |         ({const, ...} : constr_spec) =
 | 
| 41802 | 1833 | let | 
| 1834 | fun conjunct_for_sel r = | |
| 1835 | kk_rel_eq (kk_join (KK.Var (1, 0)) r) (kk_join (KK.Var (1, 1)) r) | |
| 1836 | val num_sels = num_sels_for_constr_type (snd const) | |
| 1837 | val triples = | |
| 1838 | map (const_triple rel_table | |
| 1839 | o binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize const) | |
| 1840 | (~1 upto num_sels - 1) | |
| 1841 | val set_r = triples |> hd |> #1 | |
| 1842 | in | |
| 1843 | if num_sels = 0 then | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1844 | [kk_lone set_r] | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1845 | else if all_values_are_needed need_vals dtype then | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1846 | [] | 
| 41802 | 1847 | else | 
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1848 | [kk_all (map (KK.DeclOne o rpair set_r o pair 1) [0, 1]) | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1849 | (kk_implies | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1850 | (fold1 kk_and (map (conjunct_for_sel o #1) (tl triples))) | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1851 | (kk_rel_eq (KK.Var (1, 0)) (KK.Var (1, 1))))] | 
| 41802 | 1852 | end | 
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1853 | fun uniqueness_axioms_for_datatype hol_ctxt binarize kk need_vals rel_table | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1854 |                                    (dtype as {constrs, ...}) =
 | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1855 | maps (uniqueness_axioms_for_constr hol_ctxt binarize kk need_vals rel_table | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1856 | dtype) constrs | 
| 41802 | 1857 | |
| 1858 | fun effective_constr_max ({delta, epsilon, ...} : constr_spec) = epsilon - delta
 | |
| 1859 | fun partition_axioms_for_datatype j0 (kk as {kk_rel_eq, kk_union, ...})
 | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1860 |         need_vals rel_table (dtype as {card, constrs, ...}) =
 | 
| 41802 | 1861 | if forall #exclusive constrs then | 
| 1862 | [Integer.sum (map effective_constr_max constrs) = card |> formula_for_bool] | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1863 | else if all_values_are_needed need_vals dtype then | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1864 | [] | 
| 41802 | 1865 | else | 
| 1866 | let val rs = map (discr_rel_expr rel_table o #const) constrs in | |
| 1867 | [kk_rel_eq (fold1 kk_union rs) (KK.AtomSeq (card, j0)), | |
| 1868 | kk_disjoint_sets kk rs] | |
| 1869 | end | |
| 1870 | ||
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1871 | fun other_axioms_for_datatype _ _ _ _ _ _ _ {deep = false, ...} = []
 | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1872 | | other_axioms_for_datatype hol_ctxt binarize need_vals bits ofs kk rel_table | 
| 41802 | 1873 |                               (dtype as {typ, ...}) =
 | 
| 1874 | let val j0 = offset_of_type ofs typ in | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1875 | sel_axioms_for_datatype hol_ctxt binarize bits j0 kk need_vals rel_table | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1876 | dtype @ | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1877 | uniqueness_axioms_for_datatype hol_ctxt binarize kk need_vals rel_table | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1878 | dtype @ | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1879 | partition_axioms_for_datatype j0 kk need_vals rel_table dtype | 
| 41802 | 1880 | end | 
| 1881 | ||
| 41995 | 1882 | fun declarative_axioms_for_datatypes hol_ctxt binarize need_us need_vals | 
| 41802 | 1883 | datatype_sym_break bits ofs kk rel_table dtypes = | 
| 1884 | let | |
| 1885 | val nfas = | |
| 1886 | dtypes |> map_filter (nfa_entry_for_datatype hol_ctxt binarize kk | |
| 1887 | rel_table dtypes) | |
| 1888 | |> strongly_connected_sub_nfas | |
| 1889 | in | |
| 1890 | acyclicity_axioms_for_datatypes kk nfas @ | |
| 42001 
614ff13dc5d2
preencode value of "need" selectors in Kodkod bounds as an optimization
 blanchet parents: 
42000diff
changeset | 1891 | maps (needed_value_axioms_for_datatype ofs kk dtypes) need_vals @ | 
| 41875 | 1892 | sym_break_axioms_for_datatypes hol_ctxt binarize need_us datatype_sym_break | 
| 1893 | kk rel_table nfas dtypes @ | |
| 41998 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1894 | maps (other_axioms_for_datatype hol_ctxt binarize need_vals bits ofs kk | 
| 
c2e1503fad8f
optimize Kodkod axioms further w.r.t. "need" option
 blanchet parents: 
41997diff
changeset | 1895 | rel_table) dtypes | 
| 41802 | 1896 | end | 
| 1897 | ||
| 33192 | 1898 | end; |