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