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