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