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