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