src/HOL/Tools/Nitpick/nitpick_peephole.ML
author blanchet
Sun, 25 Apr 2010 00:25:44 +0200
changeset 36390 eee4ee6a5cbe
parent 36385 ff5f88702590
child 38126 8031d099379a
permissions -rw-r--r--
remove "show_skolems" option and change style of record declarations
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33982
1ae222745c4a fixed paths in Nitpick's ML file headers
blanchet
parents: 33705
diff changeset
     1
(*  Title:      HOL/Tools/Nitpick/nitpick_peephole.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
Peephole optimizer for Nitpick.
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_PEEPHOLE =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     9
sig
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    10
  type n_ary_index = Kodkod.n_ary_index
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    11
  type formula = Kodkod.formula
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    12
  type int_expr = Kodkod.int_expr
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    13
  type rel_expr = Kodkod.rel_expr
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    14
  type decl = Kodkod.decl
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    15
  type expr_assign = Kodkod.expr_assign
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    16
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    17
  type name_pool =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    18
    {rels: n_ary_index list,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    19
     vars: n_ary_index list,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    20
     formula_reg: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    21
     rel_reg: int}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    22
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    23
  val initial_pool : name_pool
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    24
  val not3_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    25
  val suc_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    26
  val unsigned_bit_word_sel_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    27
  val signed_bit_word_sel_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    28
  val nat_add_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    29
  val int_add_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    30
  val nat_subtract_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    31
  val int_subtract_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    32
  val nat_multiply_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    33
  val int_multiply_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    34
  val nat_divide_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    35
  val int_divide_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    36
  val nat_less_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    37
  val int_less_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    38
  val gcd_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    39
  val lcm_rel : n_ary_index
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    40
  val norm_frac_rel : n_ary_index
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    41
  val atom_for_bool : int -> bool -> rel_expr
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    42
  val formula_for_bool : bool -> formula
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    43
  val atom_for_nat : int * int -> int -> int
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    44
  val min_int_for_card : int -> int
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    45
  val max_int_for_card : int -> int
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    46
  val int_for_atom : int * int -> int -> int
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    47
  val atom_for_int : int * int -> int -> int
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
    48
  val is_twos_complement_representable : int -> int -> bool
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    49
  val inline_rel_expr : rel_expr -> bool
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    50
  val empty_n_ary_rel : int -> rel_expr
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    51
  val num_seq : int -> int -> int_expr list
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    52
  val s_and : formula -> formula -> formula
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    53
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    54
  type kodkod_constrs =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    55
    {kk_all: decl list -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    56
     kk_exist: decl list -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    57
     kk_formula_let: expr_assign list -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    58
     kk_formula_if: formula -> formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    59
     kk_or: formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    60
     kk_not: formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    61
     kk_iff: formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    62
     kk_implies: formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    63
     kk_and: formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    64
     kk_subset: rel_expr -> rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    65
     kk_rel_eq: rel_expr -> rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    66
     kk_no: rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    67
     kk_lone: rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    68
     kk_one: rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    69
     kk_some: rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    70
     kk_rel_let: expr_assign list -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    71
     kk_rel_if: formula -> rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    72
     kk_union: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    73
     kk_difference: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    74
     kk_override: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    75
     kk_intersect: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    76
     kk_product: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    77
     kk_join: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    78
     kk_closure: rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    79
     kk_reflexive_closure: rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    80
     kk_comprehension: decl list -> formula -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    81
     kk_project: rel_expr -> int_expr list -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    82
     kk_project_seq: rel_expr -> int -> int -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    83
     kk_not3: rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    84
     kk_nat_less: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    85
     kk_int_less: rel_expr -> rel_expr -> rel_expr}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    86
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    87
  val kodkod_constrs : bool -> int -> int -> int -> kodkod_constrs
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    88
end;
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    89
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    90
structure Nitpick_Peephole : NITPICK_PEEPHOLE =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    91
struct
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    92
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    93
open Kodkod
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    94
open Nitpick_Util
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    95
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    96
type name_pool =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    97
  {rels: n_ary_index list,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    98
   vars: n_ary_index list,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
    99
   formula_reg: int,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   100
   rel_reg: int}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   101
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   102
(* If you add new built-in relations, make sure to increment the counters here
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   103
   as well to avoid name clashes (which fortunately would be detected by
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   104
   Kodkodi). *)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   105
val initial_pool =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   106
  {rels = [(2, 10), (3, 20), (4, 10)], vars = [], formula_reg = 10,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   107
   rel_reg = 10}
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   108
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   109
val not3_rel = (2, 0)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   110
val suc_rel = (2, 1)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   111
val unsigned_bit_word_sel_rel = (2, 2)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   112
val signed_bit_word_sel_rel = (2, 3)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   113
val nat_add_rel = (3, 0)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   114
val int_add_rel = (3, 1)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   115
val nat_subtract_rel = (3, 2)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   116
val int_subtract_rel = (3, 3)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   117
val nat_multiply_rel = (3, 4)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   118
val int_multiply_rel = (3, 5)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   119
val nat_divide_rel = (3, 6)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   120
val int_divide_rel = (3, 7)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   121
val nat_less_rel = (3, 8)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   122
val int_less_rel = (3, 9)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   123
val gcd_rel = (3, 10)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   124
val lcm_rel = (3, 11)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   125
val norm_frac_rel = (4, 0)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   126
35385
29f81babefd7 improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents: 35284
diff changeset
   127
fun atom_for_bool j0 = Atom o Integer.add j0 o int_from_bool
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   128
fun formula_for_bool b = if b then True else False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   129
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   130
fun atom_for_nat (k, j0) n = if n < 0 orelse n >= k then ~1 else n + j0
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   131
fun min_int_for_card k = ~k div 2 + 1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   132
fun max_int_for_card k = k div 2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   133
fun int_for_atom (k, j0) j =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   134
  let val j = j - j0 in if j <= max_int_for_card k then j else j - k end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   135
fun atom_for_int (k, j0) n =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   136
  if n < min_int_for_card k orelse n > max_int_for_card k then ~1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   137
  else if n < 0 then n + k + j0
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   138
  else n + j0
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   139
fun is_twos_complement_representable bits n =
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   140
  let val max = reasonable_power 2 bits in n >= ~ max andalso n < max 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
fun is_none_product (Product (r1, r2)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   143
    is_none_product r1 orelse is_none_product r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   144
  | is_none_product None = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   145
  | is_none_product _ = false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   146
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   147
fun is_one_rel_expr (Atom _) = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   148
  | is_one_rel_expr (AtomSeq (1, _)) = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   149
  | is_one_rel_expr (Var _) = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   150
  | is_one_rel_expr _ = false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   151
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   152
fun inline_rel_expr (Product (r1, r2)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   153
    inline_rel_expr r1 andalso inline_rel_expr r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   154
  | inline_rel_expr Iden = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   155
  | inline_rel_expr Ints = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   156
  | inline_rel_expr None = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   157
  | inline_rel_expr Univ = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   158
  | inline_rel_expr (Atom _) = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   159
  | inline_rel_expr (AtomSeq _) = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   160
  | inline_rel_expr (Rel _) = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   161
  | inline_rel_expr (Var _) = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   162
  | inline_rel_expr (RelReg _) = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   163
  | inline_rel_expr _ = false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   164
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   165
fun rel_expr_equal None (Atom _) = SOME false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   166
  | rel_expr_equal None (AtomSeq (k, _)) = SOME (k = 0)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   167
  | rel_expr_equal (Atom _) None = SOME false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   168
  | rel_expr_equal (AtomSeq (k, _)) None = SOME (k = 0)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   169
  | rel_expr_equal (Atom j1) (Atom j2) = SOME (j1 = j2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   170
  | rel_expr_equal (Atom j) (AtomSeq (k, j0)) = SOME (j = j0 andalso k = 1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   171
  | rel_expr_equal (AtomSeq (k, j0)) (Atom j) = SOME (j = j0 andalso k = 1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   172
  | rel_expr_equal (AtomSeq x1) (AtomSeq x2) = SOME (x1 = x2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   173
  | rel_expr_equal r1 r2 = if r1 = r2 then SOME true else NONE
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   174
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   175
fun rel_expr_intersects (Atom j1) (Atom j2) = SOME (j1 = j2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   176
  | rel_expr_intersects (Atom j) (AtomSeq (k, j0)) = SOME (j < j0 + k)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   177
  | rel_expr_intersects (AtomSeq (k, j0)) (Atom j) = SOME (j < j0 + k)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   178
  | rel_expr_intersects (AtomSeq (k1, j01)) (AtomSeq (k2, j02)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   179
    SOME (k1 > 0 andalso k2 > 0 andalso j01 + k1 > j02 andalso j02 + k2 > j01)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   180
  | rel_expr_intersects r1 r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   181
    if is_none_product r1 orelse is_none_product r2 then SOME false else NONE
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   182
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   183
fun empty_n_ary_rel 0 = raise ARG ("Nitpick_Peephole.empty_n_ary_rel", "0")
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   184
  | empty_n_ary_rel n = funpow (n - 1) (curry Product None) None
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   185
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   186
fun decl_one_set (DeclOne (_, r)) = r
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   187
  | decl_one_set _ =
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   188
    raise ARG ("Nitpick_Peephole.decl_one_set", "not \"DeclOne\"")
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   189
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   190
fun is_Num (Num _) = true
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   191
  | is_Num _ = false
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   192
fun dest_Num (Num k) = k
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
   193
  | dest_Num _ = raise ARG ("Nitpick_Peephole.dest_Num", "not \"Num\"")
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   194
fun num_seq j0 n = map Num (index_seq j0 n)
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 occurs_in_union r (Union (r1, r2)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   197
    occurs_in_union r r1 orelse occurs_in_union r r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   198
  | occurs_in_union r r' = (r = r')
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   199
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   200
fun s_and True f2 = f2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   201
  | s_and False _ = False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   202
  | s_and f1 True = f1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   203
  | s_and _ False = False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   204
  | s_and f1 f2 = And (f1, f2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   205
36390
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   206
type kodkod_constrs =
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   207
  {kk_all: decl list -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   208
   kk_exist: decl list -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   209
   kk_formula_let: expr_assign list -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   210
   kk_formula_if: formula -> formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   211
   kk_or: formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   212
   kk_not: formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   213
   kk_iff: formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   214
   kk_implies: formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   215
   kk_and: formula -> formula -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   216
   kk_subset: rel_expr -> rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   217
   kk_rel_eq: rel_expr -> rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   218
   kk_no: rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   219
   kk_lone: rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   220
   kk_one: rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   221
   kk_some: rel_expr -> formula,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   222
   kk_rel_let: expr_assign list -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   223
   kk_rel_if: formula -> rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   224
   kk_union: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   225
   kk_difference: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   226
   kk_override: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   227
   kk_intersect: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   228
   kk_product: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   229
   kk_join: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   230
   kk_closure: rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   231
   kk_reflexive_closure: rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   232
   kk_comprehension: decl list -> formula -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   233
   kk_project: rel_expr -> int_expr list -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   234
   kk_project_seq: rel_expr -> int -> int -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   235
   kk_not3: rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   236
   kk_nat_less: rel_expr -> rel_expr -> rel_expr,
eee4ee6a5cbe remove "show_skolems" option and change style of record declarations
blanchet
parents: 36385
diff changeset
   237
   kk_int_less: rel_expr -> rel_expr -> rel_expr}
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   238
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   239
(* We assume throughout that Kodkod variables have a "one" constraint. This is
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   240
   always the case if Kodkod's skolemization is disabled. *)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   241
fun kodkod_constrs optim nat_card int_card main_j0 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   242
  let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   243
    val from_bool = atom_for_bool main_j0
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   244
    fun from_nat n = Atom (n + main_j0)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   245
    fun to_nat j = j - main_j0
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   246
    val to_int = int_for_atom (int_card, main_j0)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   247
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   248
    fun s_all _ True = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   249
      | s_all _ False = False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   250
      | s_all [] f = f
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   251
      | s_all ds (All (ds', f)) = All (ds @ ds', f)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   252
      | s_all ds f = All (ds, f)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   253
    fun s_exist _ True = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   254
      | s_exist _ False = False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   255
      | s_exist [] f = f
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   256
      | s_exist ds (Exist (ds', f)) = Exist (ds @ ds', f)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   257
      | s_exist ds f = Exist (ds, f)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   258
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   259
    fun s_formula_let _ True = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   260
      | s_formula_let _ False = False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   261
      | s_formula_let assigns f = FormulaLet (assigns, f)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   262
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   263
    fun s_not True = False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   264
      | s_not False = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   265
      | s_not (All (ds, f)) = Exist (ds, s_not f)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   266
      | s_not (Exist (ds, f)) = All (ds, s_not f)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   267
      | s_not (Or (f1, f2)) = And (s_not f1, s_not f2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   268
      | s_not (Implies (f1, f2)) = And (f1, s_not f2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   269
      | s_not (And (f1, f2)) = Or (s_not f1, s_not f2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   270
      | s_not (Not f) = f
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   271
      | s_not (No r) = Some r
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   272
      | s_not (Some r) = No r
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   273
      | s_not f = Not f
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   274
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   275
    fun s_or True _ = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   276
      | s_or False f2 = f2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   277
      | s_or _ True = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   278
      | s_or f1 False = f1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   279
      | s_or f1 f2 = if f1 = f2 then f1 else Or (f1, f2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   280
    fun s_iff True f2 = f2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   281
      | s_iff False f2 = s_not f2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   282
      | s_iff f1 True = f1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   283
      | s_iff f1 False = s_not f1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   284
      | s_iff f1 f2 = if f1 = f2 then True else Iff (f1, f2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   285
    fun s_implies True f2 = f2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   286
      | s_implies False _ = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   287
      | s_implies _ True = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   288
      | s_implies f1 False = s_not f1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   289
      | s_implies f1 f2 = if f1 = f2 then True else Implies (f1, f2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   290
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   291
    fun s_formula_if True f2 _ = f2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   292
      | s_formula_if False _ f3 = f3
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   293
      | s_formula_if f1 True f3 = s_or f1 f3
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   294
      | s_formula_if f1 False f3 = s_and (s_not f1) f3
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   295
      | s_formula_if f1 f2 True = s_implies f1 f2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   296
      | s_formula_if f1 f2 False = s_and f1 f2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   297
      | s_formula_if f f1 f2 = FormulaIf (f, f1, f2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   298
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   299
    fun s_project r is =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   300
      (case r of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   301
         Project (r1, is') =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   302
         if forall is_Num is then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   303
           s_project r1 (map (nth is' o dest_Num) is)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   304
         else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   305
           raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   306
       | _ => raise SAME ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   307
      handle SAME () =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   308
             let val n = length is in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   309
               if arity_of_rel_expr r = n andalso is = num_seq 0 n then r
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   310
               else Project (r, is)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   311
             end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   312
35284
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   313
    fun s_xone xone r =
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   314
      if is_one_rel_expr r then
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   315
        True
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   316
      else case arity_of_rel_expr r of
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   317
        1 => xone r
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   318
      | arity => foldl1 And (map (xone o s_project r o single o Num)
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   319
                                 (index_seq 0 arity))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   320
    fun s_no None = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   321
      | s_no (Product (r1, r2)) = s_or (s_no r1) (s_no r2)
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   322
      | s_no (Intersect (Closure (Rel x), Iden)) = Acyclic x
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   323
      | s_no r = if is_one_rel_expr r then False else No r
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   324
    fun s_lone None = True
35284
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   325
      | s_lone r = s_xone Lone r
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   326
    fun s_one None = False
35284
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   327
      | s_one r = s_xone One r
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   328
    fun s_some None = False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   329
      | s_some (Atom _) = True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   330
      | s_some (Product (r1, r2)) = s_and (s_some r1) (s_some r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   331
      | s_some r = if is_one_rel_expr r then True else Some r
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   332
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   333
    fun s_not3 (Atom j) = Atom (if j = main_j0 then j + 1 else j - 1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   334
      | s_not3 (r as Join (r1, r2)) =
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   335
        if r2 = Rel not3_rel then r1 else Join (r, Rel not3_rel)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   336
      | s_not3 r = Join (r, Rel not3_rel)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   337
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   338
    fun s_rel_eq r1 r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   339
      (case (r1, r2) of
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   340
         (Join (r11, Rel x), _) =>
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   341
         if x = not3_rel then s_rel_eq r11 (s_not3 r2) else raise SAME ()
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   342
       | (_, Join (r21, Rel x)) =>
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   343
         if x = not3_rel then s_rel_eq r21 (s_not3 r1) else raise SAME ()
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   344
       | (RelIf (f, r11, r12), _) =>
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   345
         if inline_rel_expr r2 then
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   346
           s_formula_if f (s_rel_eq r11 r2) (s_rel_eq r12 r2)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   347
         else
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   348
           raise SAME ()
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   349
       | (_, RelIf (f, r21, r22)) =>
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   350
         if inline_rel_expr r1 then
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   351
           s_formula_if f (s_rel_eq r1 r21) (s_rel_eq r1 r22)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   352
         else
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   353
           raise SAME ()
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   354
       | (RelLet (bs, r1'), Atom _) => s_formula_let bs (s_rel_eq r1' r2)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   355
       | (Atom _, RelLet (bs, r2')) => s_formula_let bs (s_rel_eq r1 r2')
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   356
       | _ => raise SAME ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   357
      handle SAME () =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   358
             case rel_expr_equal r1 r2 of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   359
               SOME true => True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   360
             | SOME false => False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   361
             | NONE =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   362
               case (r1, r2) of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   363
                 (_, RelIf (f, r21, r22)) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   364
                  if inline_rel_expr r1 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   365
                    s_formula_if f (s_rel_eq r1 r21) (s_rel_eq r1 r22)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   366
                  else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   367
                    RelEq (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   368
               | (RelIf (f, r11, r12), _) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   369
                  if inline_rel_expr r2 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   370
                    s_formula_if f (s_rel_eq r11 r2) (s_rel_eq r12 r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   371
                  else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   372
                    RelEq (r1, r2)
34126
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   373
               | (_, None) => s_no r1
8a2c5d7aff51 polished Nitpick's binary integer support etc.;
blanchet
parents: 34124
diff changeset
   374
               | (None, _) => s_no r2
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   375
               | _ => RelEq (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   376
    fun s_subset (Atom j1) (Atom j2) = formula_for_bool (j1 = j2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   377
      | s_subset (Atom j) (AtomSeq (k, j0)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   378
        formula_for_bool (j >= j0 andalso j < j0 + k)
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 34982
diff changeset
   379
      | s_subset (Union (r11, r12)) r2 =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   380
        s_and (s_subset r11 r2) (s_subset r12 r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   381
      | s_subset r1 (r2 as Union (r21, r22)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   382
        if is_one_rel_expr r1 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   383
          s_or (s_subset r1 r21) (s_subset r1 r22)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   384
        else
34936
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34126
diff changeset
   385
          if s_subset r1 r21 = True orelse s_subset r1 r22 = True orelse
c4f04bee79f3 some work on Nitpick's support for quotient types;
blanchet
parents: 34126
diff changeset
   386
             r1 = r2 then
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   387
            True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   388
          else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   389
            Subset (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   390
      | s_subset r1 r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   391
        if r1 = r2 orelse is_none_product r1 then True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   392
        else if is_none_product r2 then s_no r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   393
        else if forall is_one_rel_expr [r1, r2] then s_rel_eq r1 r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   394
        else Subset (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   395
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   396
    fun s_rel_let [b as AssignRelReg (x', r')] (r as RelReg x) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   397
        if x = x' then r' else RelLet ([b], r)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   398
      | s_rel_let bs r = RelLet (bs, r)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   399
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   400
    fun s_rel_if f r1 r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   401
      (case (f, r1, r2) of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   402
         (True, _, _) => r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   403
       | (False, _, _) => r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   404
       | (No r1', None, RelIf (One r2', r3', r4')) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   405
         if r1' = r2' andalso r2' = r3' then s_rel_if (Lone r1') r1' r4'
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   406
         else raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   407
       | _ => raise SAME ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   408
      handle SAME () => if r1 = r2 then r1 else RelIf (f, r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   409
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   410
    fun s_union r1 (Union (r21, r22)) = s_union (s_union r1 r21) r22
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   411
      | s_union r1 r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   412
        if is_none_product r1 then r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   413
        else if is_none_product r2 then r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   414
        else if r1 = r2 then r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   415
        else if occurs_in_union r2 r1 then r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   416
        else Union (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   417
    fun s_difference r1 r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   418
      if is_none_product r1 orelse is_none_product r2 then r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   419
      else if r1 = r2 then empty_n_ary_rel (arity_of_rel_expr r1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   420
      else Difference (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   421
    fun s_override r1 r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   422
      if is_none_product r2 then r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   423
      else if is_none_product r1 then r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   424
      else Override (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   425
    fun s_intersect r1 r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   426
      case rel_expr_intersects r1 r2 of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   427
        SOME true => if r1 = r2 then r1 else Intersect (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   428
      | SOME false => empty_n_ary_rel (arity_of_rel_expr r1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   429
      | NONE => if is_none_product r1 then r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   430
                else if is_none_product r2 then r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   431
                else Intersect (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   432
    fun s_product r1 r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   433
      if is_none_product r1 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   434
        Product (r1, empty_n_ary_rel (arity_of_rel_expr r2))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   435
      else if is_none_product r2 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   436
        Product (empty_n_ary_rel (arity_of_rel_expr r1), r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   437
      else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   438
        Product (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   439
    fun s_join r1 (Product (Product (r211, r212), r22)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   440
        Product (s_join r1 (Product (r211, r212)), r22)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   441
      | s_join (Product (r11, Product (r121, r122))) r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   442
        Product (r11, s_join (Product (r121, r122)) r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   443
      | s_join None r = empty_n_ary_rel (arity_of_rel_expr r - 1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   444
      | s_join r None = empty_n_ary_rel (arity_of_rel_expr r - 1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   445
      | s_join (Product (None, None)) r = empty_n_ary_rel (arity_of_rel_expr r)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   446
      | s_join r (Product (None, None)) = empty_n_ary_rel (arity_of_rel_expr r)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   447
      | s_join Iden r2 = r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   448
      | s_join r1 Iden = r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   449
      | s_join (Product (r1, r2)) Univ =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   450
        if arity_of_rel_expr r2 = 1 then r1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   451
        else Product (r1, s_join r2 Univ)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   452
      | s_join Univ (Product (r1, r2)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   453
        if arity_of_rel_expr r1 = 1 then r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   454
        else Product (s_join Univ r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   455
      | s_join r1 (r2 as Product (r21, r22)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   456
        if arity_of_rel_expr r1 = 1 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   457
          case rel_expr_intersects r1 r21 of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   458
            SOME true => r22
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   459
          | SOME false => empty_n_ary_rel (arity_of_rel_expr r2 - 1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   460
          | NONE => Join (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   461
        else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   462
          Join (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   463
      | s_join (r1 as Product (r11, r12)) r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   464
        if arity_of_rel_expr r2 = 1 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   465
          case rel_expr_intersects r2 r12 of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   466
            SOME true => r11
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   467
          | SOME false => empty_n_ary_rel (arity_of_rel_expr r1 - 1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   468
          | NONE => Join (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   469
        else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   470
          Join (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   471
      | s_join r1 (r2 as RelIf (f, r21, r22)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   472
        if inline_rel_expr r1 then s_rel_if f (s_join r1 r21) (s_join r1 r22)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   473
        else Join (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   474
      | s_join (r1 as RelIf (f, r11, r12)) r2 =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   475
        if inline_rel_expr r2 then s_rel_if f (s_join r11 r2) (s_join r12 r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   476
        else Join (r1, r2)
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 34982
diff changeset
   477
      | s_join (r1 as Atom j1) (r2 as Rel (x as (2, _))) =
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   478
        if x = suc_rel then
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   479
          let val n = to_nat j1 + 1 in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   480
            if n < nat_card then from_nat n else None
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   481
          end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   482
        else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   483
          Join (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   484
      | s_join r1 (r2 as Project (r21, Num k :: is)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   485
        if k = arity_of_rel_expr r21 - 1 andalso arity_of_rel_expr r1 = 1 then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   486
          s_project (s_join r21 r1) is
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   487
        else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   488
          Join (r1, r2)
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 34982
diff changeset
   489
      | s_join r1 (Join (r21, r22 as Rel (x as (3, _)))) =
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   490
        ((if x = nat_add_rel then
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   491
            case (r21, r1) of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   492
              (Atom j1, Atom j2) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   493
              let val n = to_nat j1 + to_nat j2 in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   494
                if n < nat_card then from_nat n else None
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   495
              end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   496
            | (Atom j, r) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   497
              (case to_nat j of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   498
                 0 => r
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   499
               | 1 => s_join r (Rel suc_rel)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   500
               | _ => raise SAME ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   501
            | (r, Atom j) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   502
              (case to_nat j of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   503
                 0 => r
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   504
               | 1 => s_join r (Rel suc_rel)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   505
               | _ => raise SAME ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   506
            | _ => raise SAME ()
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   507
          else if x = nat_subtract_rel then
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   508
            case (r21, r1) of
33705
947184dc75c9 removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents: 33232
diff changeset
   509
              (Atom j1, Atom j2) => from_nat (nat_minus (to_nat j1) (to_nat j2))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   510
            | _ => raise SAME ()
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   511
          else if x = nat_multiply_rel then
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   512
            case (r21, r1) of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   513
              (Atom j1, Atom j2) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   514
              let val n = to_nat j1 * to_nat j2 in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   515
                if n < nat_card then from_nat n else None
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   516
              end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   517
            | (Atom j, r) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   518
              (case to_nat j of 0 => Atom j | 1 => r | _ => raise SAME ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   519
            | (r, Atom j) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   520
              (case to_nat j of 0 => Atom j | 1 => r | _ => raise SAME ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   521
            | _ => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   522
          else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   523
            raise SAME ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   524
         handle SAME () => List.foldr Join r22 [r1, r21])
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   525
      | s_join r1 r2 = Join (r1, r2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   526
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   527
    fun s_closure Iden = Iden
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   528
      | s_closure r = if is_none_product r then r else Closure r
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   529
    fun s_reflexive_closure Iden = Iden
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   530
      | s_reflexive_closure r =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   531
        if is_none_product r then Iden else ReflexiveClosure r
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   532
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   533
    fun s_comprehension ds False = empty_n_ary_rel (length ds)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   534
      | s_comprehension ds True = fold1 s_product (map decl_one_set ds)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   535
      | s_comprehension [d as DeclOne ((1, j1), r)]
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   536
                        (f as RelEq (Var (1, j2), Atom j)) =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   537
        if j1 = j2 andalso rel_expr_intersects (Atom j) r = SOME true then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   538
          Atom j
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   539
        else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   540
          Comprehension ([d], f)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   541
      | s_comprehension ds f = Comprehension (ds, f)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   542
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   543
    fun s_project_seq r =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   544
      let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   545
        fun aux arity r j0 n =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   546
          if j0 = 0 andalso arity = n then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   547
            r
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   548
          else case r of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   549
            RelIf (f, r1, r2) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   550
            s_rel_if f (aux arity r1 j0 n) (aux arity r2 j0 n)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   551
          | Product (r1, r2) =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   552
            let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   553
              val arity2 = arity_of_rel_expr r2
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   554
              val arity1 = arity - arity2
33705
947184dc75c9 removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents: 33232
diff changeset
   555
              val n1 = Int.min (nat_minus arity1 j0, n)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   556
              val n2 = n - n1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   557
              fun one () = aux arity1 r1 j0 n1
33705
947184dc75c9 removed a few global names in Nitpick (styp, nat_less, pairf)
blanchet
parents: 33232
diff changeset
   558
              fun two () = aux arity2 r2 (nat_minus j0 arity1) n2
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   559
            in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   560
              case (n1, n2) of
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   561
                (0, _) => s_rel_if (s_some r1) (two ()) (empty_n_ary_rel n2)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   562
              | (_, 0) => s_rel_if (s_some r2) (one ()) (empty_n_ary_rel n1)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   563
              | _ => s_product (one ()) (two ())
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   564
            end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   565
          | _ => s_project r (num_seq j0 n)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   566
      in aux (arity_of_rel_expr r) r end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   567
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   568
    fun s_nat_less (Atom j1) (Atom j2) = from_bool (j1 < j2)
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   569
      | s_nat_less r1 r2 = fold s_join [r1, r2] (Rel nat_less_rel)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   570
    fun s_int_less (Atom j1) (Atom j2) = from_bool (to_int j1 < to_int j2)
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   571
      | s_int_less r1 r2 = fold s_join [r1, r2] (Rel int_less_rel)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   572
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   573
    fun d_project_seq r j0 n = Project (r, num_seq j0 n)
34124
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   574
    fun d_not3 r = Join (r, Rel not3_rel)
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   575
    fun d_nat_less r1 r2 = List.foldl Join (Rel nat_less_rel) [r1, r2]
c4628a1dcf75 added support for binary nat/int representation to Nitpick
blanchet
parents: 33982
diff changeset
   576
    fun d_int_less r1 r2 = List.foldl Join (Rel int_less_rel) [r1, r2]
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   577
  in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   578
    if optim then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   579
      {kk_all = s_all, kk_exist = s_exist, kk_formula_let = s_formula_let,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   580
       kk_formula_if = s_formula_if, kk_or = s_or, kk_not = s_not,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   581
       kk_iff = s_iff, kk_implies = s_implies, kk_and = s_and,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   582
       kk_subset = s_subset, kk_rel_eq = s_rel_eq, kk_no = s_no,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   583
       kk_lone = s_lone, kk_one = s_one, kk_some = s_some,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   584
       kk_rel_let = s_rel_let, kk_rel_if = s_rel_if, kk_union = s_union,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   585
       kk_difference = s_difference, kk_override = s_override,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   586
       kk_intersect = s_intersect, kk_product = s_product, kk_join = s_join,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   587
       kk_closure = s_closure, kk_reflexive_closure = s_reflexive_closure,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   588
       kk_comprehension = s_comprehension, kk_project = s_project,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   589
       kk_project_seq = s_project_seq, kk_not3 = s_not3,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   590
       kk_nat_less = s_nat_less, kk_int_less = s_int_less}
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   591
    else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   592
      {kk_all = curry All, kk_exist = curry Exist,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   593
       kk_formula_let = curry FormulaLet, kk_formula_if = curry3 FormulaIf,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   594
       kk_or = curry Or,kk_not = Not, kk_iff = curry Iff, kk_implies = curry
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   595
       Implies, kk_and = curry And, kk_subset = curry Subset, kk_rel_eq = curry
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   596
       RelEq, kk_no = No, kk_lone = Lone, kk_one = One, kk_some = Some,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   597
       kk_rel_let = curry RelLet, kk_rel_if = curry3 RelIf, kk_union = curry
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   598
       Union, kk_difference = curry Difference, kk_override = curry Override,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   599
       kk_intersect = curry Intersect, kk_product = curry Product,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   600
       kk_join = curry Join, kk_closure = Closure,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   601
       kk_reflexive_closure = ReflexiveClosure, kk_comprehension = curry
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   602
       Comprehension, kk_project = curry Project,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   603
       kk_project_seq = d_project_seq, kk_not3 = d_not3,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   604
       kk_nat_less = d_nat_less, kk_int_less = d_int_less}
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   605
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   606
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   607
end;