src/HOL/Tools/Nitpick/minipick.ML
author blanchet
Wed, 04 Aug 2010 10:39:35 +0200
changeset 38190 b02e204b613a
parent 37678 0040bafffdef
child 38516 307669429dc1
permissions -rw-r--r--
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
     1
(*  Title:      HOL/Tools/Nitpick/minipick.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: 34124
diff changeset
     3
    Copyright   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
Finite model generation for HOL formulas using Kodkod, minimalistic version.
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 MINIPICK =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
     9
sig
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    10
  datatype rep = SRep | RRep
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    11
  type styp = Nitpick_Util.styp
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    12
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    13
  val vars_for_bound_var :
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    14
    (typ -> int) -> rep -> typ list -> int -> Kodkod.rel_expr list
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    15
  val rel_expr_for_bound_var :
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    16
    (typ -> int) -> rep -> typ list -> int -> Kodkod.rel_expr
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    17
  val decls_for : rep -> (typ -> int) -> typ list -> typ -> Kodkod.decl list
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    18
  val false_atom : Kodkod.rel_expr
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    19
  val true_atom : Kodkod.rel_expr
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    20
  val formula_from_atom : Kodkod.rel_expr -> Kodkod.formula
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    21
  val atom_from_formula : Kodkod.formula -> Kodkod.rel_expr
35284
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
    22
  val kodkod_problem_from_term :
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
    23
    Proof.context -> (typ -> int) -> term -> Kodkod.problem
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
    24
  val solve_any_kodkod_problem : theory -> Kodkod.problem list -> string
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    25
end;
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    26
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    27
structure Minipick : MINIPICK =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    28
struct
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    29
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    30
open Kodkod
33232
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    31
open Nitpick_Util
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    32
open Nitpick_HOL
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    33
open Nitpick_Peephole
f93390060bbe internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents: 33192
diff changeset
    34
open Nitpick_Kodkod
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    35
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    36
datatype rep = SRep | RRep
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    37
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
    38
fun check_type ctxt (Type (@{type_name fun}, Ts)) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
    39
    List.app (check_type ctxt) Ts
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 37678
diff changeset
    40
  | check_type ctxt (Type (@{type_name prod}, Ts)) =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
    41
    List.app (check_type ctxt) Ts
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    42
  | check_type _ @{typ bool} = ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    43
  | check_type _ (TFree (_, @{sort "{}"})) = ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    44
  | check_type _ (TFree (_, @{sort HOL.type})) = ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    45
  | check_type ctxt T =
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    46
    raise NOT_SUPPORTED ("type " ^ quote (Syntax.string_of_typ ctxt T))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    47
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
    48
fun atom_schema_of SRep card (Type (@{type_name fun}, [T1, T2])) =
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    49
    replicate_list (card T1) (atom_schema_of SRep card T2)
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
    50
  | atom_schema_of RRep card (Type (@{type_name fun}, [T1, @{typ bool}])) =
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    51
    atom_schema_of SRep card T1
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
    52
  | atom_schema_of RRep card (Type (@{type_name fun}, [T1, T2])) =
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    53
    atom_schema_of SRep card T1 @ atom_schema_of RRep card T2
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 37678
diff changeset
    54
  | atom_schema_of _ card (Type (@{type_name prod}, Ts)) =
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
    55
    maps (atom_schema_of SRep card) Ts
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    56
  | atom_schema_of _ card T = [card T]
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    57
val arity_of = length ooo atom_schema_of
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    58
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    59
fun index_for_bound_var _ [_] 0 = 0
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    60
  | index_for_bound_var card (_ :: Ts) 0 =
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    61
    index_for_bound_var card Ts 0 + arity_of SRep card (hd Ts)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    62
  | index_for_bound_var card Ts n = index_for_bound_var card (tl Ts) (n - 1)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    63
fun vars_for_bound_var card R Ts j =
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    64
  map (curry Var 1) (index_seq (index_for_bound_var card Ts j)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    65
                               (arity_of R card (nth Ts j)))
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    66
val rel_expr_for_bound_var = foldl1 Product oooo vars_for_bound_var
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    67
fun decls_for R card Ts T =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    68
  map2 (curry DeclOne o pair 1)
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    69
       (index_seq (index_for_bound_var card (T :: Ts) 0)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    70
                  (arity_of R card (nth (T :: Ts) 0)))
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    71
       (map (AtomSeq o rpair 0) (atom_schema_of R card T))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    72
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    73
val atom_product = foldl1 Product o map Atom
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    74
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    75
val false_atom = Atom 0
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    76
val true_atom = Atom 1
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    77
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    78
fun formula_from_atom r = RelEq (r, true_atom)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    79
fun atom_from_formula f = RelIf (f, true_atom, false_atom)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    80
35185
9b8f351cced6 added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
blanchet
parents: 35028
diff changeset
    81
fun kodkod_formula_from_term ctxt card frees =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    82
  let
37396
18a1e9c7acb0 remove needless variables
blanchet
parents: 36385
diff changeset
    83
    fun R_rep_from_S_rep (Type (@{type_name fun}, [T1, @{typ bool}])) r =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    84
        let
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    85
          val jss = atom_schema_of SRep card T1 |> map (rpair 0)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    86
                    |> all_combinations
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    87
        in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    88
          map2 (fn i => fn js =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    89
                   RelIf (formula_from_atom (Project (r, [Num i])),
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    90
                          atom_product js, empty_n_ary_rel (length js)))
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    91
               (index_seq 0 (length jss)) jss
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    92
          |> foldl1 Union
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    93
        end
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
    94
      | R_rep_from_S_rep (Type (@{type_name fun}, [T1, T2])) r =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    95
        let
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    96
          val jss = atom_schema_of SRep card T1 |> map (rpair 0)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    97
                    |> all_combinations
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
    98
          val arity2 = arity_of SRep card T2
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
    99
        in
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   100
          map2 (fn i => fn js =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   101
                   Product (atom_product js,
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   102
                            Project (r, num_seq (i * arity2) arity2)
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   103
                            |> R_rep_from_S_rep T2))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   104
               (index_seq 0 (length jss)) jss
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   105
          |> foldl1 Union
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   106
        end
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   107
      | R_rep_from_S_rep _ r = r
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   108
    fun S_rep_from_R_rep Ts (T as Type (@{type_name fun}, _)) r =
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   109
        Comprehension (decls_for SRep card Ts T,
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   110
            RelEq (R_rep_from_S_rep T
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   111
                       (rel_expr_for_bound_var card SRep (T :: Ts) 0), r))
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   112
      | S_rep_from_R_rep _ _ r = r
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   113
    fun to_F Ts t =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   114
      (case t of
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   115
         @{const Not} $ t1 => Not (to_F Ts t1)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   116
       | @{const False} => False
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   117
       | @{const True} => True
37396
18a1e9c7acb0 remove needless variables
blanchet
parents: 36385
diff changeset
   118
       | Const (@{const_name All}, _) $ Abs (_, T, t') =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   119
         All (decls_for SRep card Ts T, to_F (T :: Ts) t')
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   120
       | (t0 as Const (@{const_name All}, _)) $ t1 =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   121
         to_F Ts (t0 $ eta_expand Ts t1 1)
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35185
diff changeset
   122
       | Const (@{const_name Ex}, _) $ Abs (_, T, t') =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   123
         Exist (decls_for SRep card Ts T, to_F (T :: Ts) t')
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   124
       | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   125
         to_F Ts (t0 $ eta_expand Ts t1 1)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   126
       | Const (@{const_name "op ="}, _) $ t1 $ t2 =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   127
         RelEq (to_R_rep Ts t1, to_R_rep Ts t2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   128
       | Const (@{const_name ord_class.less_eq},
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   129
                Type (@{type_name fun},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   130
                      [Type (@{type_name fun}, [_, @{typ bool}]), _]))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   131
         $ t1 $ t2 =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   132
         Subset (to_R_rep Ts t1, to_R_rep Ts t2)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   133
       | @{const "op &"} $ t1 $ t2 => And (to_F Ts t1, to_F Ts t2)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   134
       | @{const "op |"} $ t1 $ t2 => Or (to_F Ts t1, to_F Ts t2)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   135
       | @{const "op -->"} $ t1 $ t2 => Implies (to_F Ts t1, to_F Ts t2)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   136
       | t1 $ t2 => Subset (to_S_rep Ts t2, to_R_rep Ts t1)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   137
       | Free _ => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   138
       | Term.Var _ => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   139
       | Bound _ => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   140
       | Const (s, _) => raise NOT_SUPPORTED ("constant " ^ quote s)
35185
9b8f351cced6 added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
blanchet
parents: 35028
diff changeset
   141
       | _ => raise TERM ("Minipick.kodkod_formula_from_term.to_F", [t]))
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   142
      handle SAME () => formula_from_atom (to_R_rep Ts t)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   143
    and to_S_rep Ts t =
36385
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   144
      case t of
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   145
        Const (@{const_name Pair}, _) $ t1 $ t2 =>
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   146
        Product (to_S_rep Ts t1, to_S_rep Ts t2)
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   147
      | Const (@{const_name Pair}, _) $ _ => to_S_rep Ts (eta_expand Ts t 1)
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   148
      | Const (@{const_name Pair}, _) => to_S_rep Ts (eta_expand Ts t 2)
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   149
      | Const (@{const_name fst}, _) $ t1 =>
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   150
        let val fst_arity = arity_of SRep card (fastype_of1 (Ts, t)) in
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   151
          Project (to_S_rep Ts t1, num_seq 0 fst_arity)
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   152
        end
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   153
      | Const (@{const_name fst}, _) => to_S_rep Ts (eta_expand Ts t 1)
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   154
      | Const (@{const_name snd}, _) $ t1 =>
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   155
        let
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   156
          val pair_arity = arity_of SRep card (fastype_of1 (Ts, t1))
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   157
          val snd_arity = arity_of SRep card (fastype_of1 (Ts, t))
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   158
          val fst_arity = pair_arity - snd_arity
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   159
        in Project (to_S_rep Ts t1, num_seq fst_arity snd_arity) end
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   160
      | Const (@{const_name snd}, _) => to_S_rep Ts (eta_expand Ts t 1)
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   161
      | Bound j => rel_expr_for_bound_var card SRep Ts j
ff5f88702590 remove type annotations as comments;
blanchet
parents: 35699
diff changeset
   162
      | _ => S_rep_from_R_rep Ts (fastype_of1 (Ts, t)) (to_R_rep Ts t)
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   163
    and to_R_rep Ts t =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   164
      (case t of
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   165
         @{const Not} => to_R_rep Ts (eta_expand Ts t 1)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   166
       | Const (@{const_name All}, _) => to_R_rep Ts (eta_expand Ts t 1)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   167
       | Const (@{const_name Ex}, _) => to_R_rep Ts (eta_expand Ts t 1)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   168
       | Const (@{const_name "op ="}, _) $ _ => to_R_rep Ts (eta_expand Ts t 1)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   169
       | Const (@{const_name "op ="}, _) => to_R_rep Ts (eta_expand Ts t 2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   170
       | Const (@{const_name ord_class.less_eq},
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   171
                Type (@{type_name fun},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   172
                      [Type (@{type_name fun}, [_, @{typ bool}]), _])) $ _ =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   173
         to_R_rep Ts (eta_expand Ts t 1)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   174
       | Const (@{const_name ord_class.less_eq}, _) =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   175
         to_R_rep Ts (eta_expand Ts t 2)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   176
       | @{const "op &"} $ _ => to_R_rep Ts (eta_expand Ts t 1)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   177
       | @{const "op &"} => to_R_rep Ts (eta_expand Ts t 2)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   178
       | @{const "op |"} $ _ => to_R_rep Ts (eta_expand Ts t 1)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   179
       | @{const "op |"} => to_R_rep Ts (eta_expand Ts t 2)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   180
       | @{const "op -->"} $ _ => to_R_rep Ts (eta_expand Ts t 1)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   181
       | @{const "op -->"} => to_R_rep Ts (eta_expand Ts t 2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   182
       | Const (@{const_name bot_class.bot},
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   183
                T as Type (@{type_name fun}, [_, @{typ bool}])) =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   184
         empty_n_ary_rel (arity_of RRep card T)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   185
       | Const (@{const_name insert}, _) $ t1 $ t2 =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   186
         Union (to_S_rep Ts t1, to_R_rep Ts t2)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   187
       | Const (@{const_name insert}, _) $ _ => to_R_rep Ts (eta_expand Ts t 1)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   188
       | Const (@{const_name insert}, _) => to_R_rep Ts (eta_expand Ts t 2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   189
       | Const (@{const_name trancl}, _) $ t1 =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   190
         if arity_of RRep card (fastype_of1 (Ts, t1)) = 2 then
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   191
           Closure (to_R_rep Ts t1)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   192
         else
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   193
           raise NOT_SUPPORTED "transitive closure for function or pair type"
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   194
       | Const (@{const_name trancl}, _) => to_R_rep Ts (eta_expand Ts t 1)
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34982
diff changeset
   195
       | Const (@{const_name semilattice_inf_class.inf},
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   196
                Type (@{type_name fun},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   197
                      [Type (@{type_name fun}, [_, @{typ bool}]), _]))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   198
         $ t1 $ t2 =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   199
         Intersect (to_R_rep Ts t1, to_R_rep Ts t2)
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34982
diff changeset
   200
       | Const (@{const_name semilattice_inf_class.inf}, _) $ _ =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   201
         to_R_rep Ts (eta_expand Ts t 1)
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34982
diff changeset
   202
       | Const (@{const_name semilattice_inf_class.inf}, _) =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   203
         to_R_rep Ts (eta_expand Ts t 2)
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34982
diff changeset
   204
       | Const (@{const_name semilattice_sup_class.sup},
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   205
                Type (@{type_name fun},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   206
                      [Type (@{type_name fun}, [_, @{typ bool}]), _]))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   207
         $ t1 $ t2 =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   208
         Union (to_R_rep Ts t1, to_R_rep Ts t2)
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34982
diff changeset
   209
       | Const (@{const_name semilattice_sup_class.sup}, _) $ _ =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   210
         to_R_rep Ts (eta_expand Ts t 1)
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34982
diff changeset
   211
       | Const (@{const_name semilattice_sup_class.sup}, _) =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   212
         to_R_rep Ts (eta_expand Ts t 2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   213
       | Const (@{const_name minus_class.minus},
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   214
                Type (@{type_name fun},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   215
                      [Type (@{type_name fun}, [_, @{typ bool}]), _]))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   216
         $ t1 $ t2 =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   217
         Difference (to_R_rep Ts t1, to_R_rep Ts t2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   218
       | Const (@{const_name minus_class.minus},
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   219
                Type (@{type_name fun},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   220
                      [Type (@{type_name fun}, [_, @{typ bool}]), _])) $ _ =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   221
         to_R_rep Ts (eta_expand Ts t 1)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   222
       | Const (@{const_name minus_class.minus},
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   223
                Type (@{type_name fun},
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   224
                      [Type (@{type_name fun}, [_, @{typ bool}]), _])) =>
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   225
         to_R_rep Ts (eta_expand Ts t 2)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   226
       | Const (@{const_name Pair}, _) $ _ $ _ => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   227
       | Const (@{const_name Pair}, _) $ _ => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   228
       | Const (@{const_name Pair}, _) => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   229
       | Const (@{const_name fst}, _) $ _ => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   230
       | Const (@{const_name fst}, _) => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   231
       | Const (@{const_name snd}, _) $ _ => raise SAME ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   232
       | Const (@{const_name snd}, _) => raise SAME ()
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   233
       | Const (_, @{typ bool}) => atom_from_formula (to_F Ts t)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   234
       | Free (x as (_, T)) =>
34121
5e831d805118 get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents: 33980
diff changeset
   235
         Rel (arity_of RRep card T, find_index (curry (op =) x) frees)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   236
       | Term.Var _ => raise NOT_SUPPORTED "schematic variables"
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35185
diff changeset
   237
       | Bound _ => raise SAME ()
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   238
       | Abs (_, T, t') =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   239
         (case fastype_of1 (T :: Ts, t') of
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   240
            @{typ bool} => Comprehension (decls_for SRep card Ts T,
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   241
                                          to_F (T :: Ts) t')
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   242
          | T' => Comprehension (decls_for SRep card Ts T @
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   243
                                 decls_for RRep card (T :: Ts) T',
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   244
                                 Subset (rel_expr_for_bound_var card RRep
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   245
                                                              (T' :: T :: Ts) 0,
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   246
                                         to_R_rep (T :: Ts) t')))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   247
       | t1 $ t2 =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   248
         (case fastype_of1 (Ts, t) of
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   249
            @{typ bool} => atom_from_formula (to_F Ts t)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   250
          | T =>
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   251
            let val T2 = fastype_of1 (Ts, t2) in
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   252
              case arity_of SRep card T2 of
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   253
                1 => Join (to_S_rep Ts t2, to_R_rep Ts t1)
35280
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35185
diff changeset
   254
              | arity2 =>
54ab4921f826 fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents: 35185
diff changeset
   255
                let val res_arity = arity_of RRep card T in
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   256
                  Project (Intersect
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   257
                      (Product (to_S_rep Ts t2,
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   258
                                atom_schema_of RRep card T
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   259
                                |> map (AtomSeq o rpair 0) |> foldl1 Product),
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   260
                       to_R_rep Ts t1),
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   261
                      num_seq arity2 res_arity)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   262
                end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   263
            end)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   264
       | _ => raise NOT_SUPPORTED ("term " ^
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   265
                                   quote (Syntax.string_of_term ctxt t)))
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   266
      handle SAME () => R_rep_from_S_rep (fastype_of1 (Ts, t)) (to_S_rep Ts t)
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   267
  in to_F [] end
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   268
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   269
fun bound_for_free card i (s, T) =
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   270
  let val js = atom_schema_of RRep card T in
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   271
    ([((length js, i), s)],
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   272
     [TupleSet [], atom_schema_of RRep card T |> map (rpair 0)
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   273
                   |> tuple_set_from_atom_schema])
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   274
  end
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   275
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   276
fun declarative_axiom_for_rel_expr card Ts (Type (@{type_name fun}, [T1, T2]))
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   277
                                   r =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   278
    if body_type T2 = bool_T then
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   279
      True
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   280
    else
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   281
      All (decls_for SRep card Ts T1,
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   282
           declarative_axiom_for_rel_expr card (T1 :: Ts) T2
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   283
               (List.foldl Join r (vars_for_bound_var card SRep (T1 :: Ts) 0)))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   284
  | declarative_axiom_for_rel_expr _ _ _ r = One r
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   285
fun declarative_axiom_for_free card i (_, T) =
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   286
  declarative_axiom_for_rel_expr card [] T (Rel (arity_of RRep card T, i))
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   287
35284
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   288
fun kodkod_problem_from_term ctxt raw_card t =
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   289
  let
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   290
    val thy = ProofContext.theory_of ctxt
35665
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   291
    fun card (Type (@{type_name fun}, [T1, T2])) =
ff2bf50505ab added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents: 35625
diff changeset
   292
        reasonable_power (card T2) (card T1)
38190
b02e204b613a get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents: 37678
diff changeset
   293
      | card (Type (@{type_name prod}, [T1, T2])) = card T1 * card T2
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   294
      | card @{typ bool} = 2
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   295
      | card T = Int.max (1, raw_card T)
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35333
diff changeset
   296
    val neg_t = @{const Not} $ Object_Logic.atomize_term thy t
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   297
    val _ = fold_types (K o check_type ctxt) neg_t ()
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   298
    val frees = Term.add_frees neg_t []
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   299
    val bounds = map2 (bound_for_free card) (index_seq 0 (length frees)) frees
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   300
    val declarative_axioms =
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   301
      map2 (declarative_axiom_for_free card) (index_seq 0 (length frees)) frees
35185
9b8f351cced6 added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
blanchet
parents: 35028
diff changeset
   302
    val formula = kodkod_formula_from_term ctxt card frees neg_t
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   303
                  |> fold_rev (curry And) declarative_axioms
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   304
    val univ_card = univ_card 0 0 0 bounds formula
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   305
  in
35284
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   306
    {comment = "", settings = [], univ_card = univ_card, tuple_assigns = [],
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   307
     bounds = bounds, int_bounds = [], expr_assigns = [], formula = formula}
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   308
  end
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   309
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   310
fun solve_any_kodkod_problem thy problems =
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   311
  let
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   312
    val {overlord, ...} = Nitpick_Isar.default_params thy []
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   313
    val max_threads = 1
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   314
    val max_solutions = 1
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   315
  in
9edc2bd6d2bd enabled Nitpick's support for quotient types + shortened the Nitpick tests a bit
blanchet
parents: 35280
diff changeset
   316
    case solve_any_problem overlord NONE max_threads max_solutions problems of
35699
9ed327529a44 improve precision of "card" in Nitpick
blanchet
parents: 35665
diff changeset
   317
      JavaNotInstalled => "unknown"
9ed327529a44 improve precision of "card" in Nitpick
blanchet
parents: 35665
diff changeset
   318
    | KodkodiNotInstalled => "unknown"
35333
f61de25f71f9 distinguish between Kodkodi warnings and errors in Nitpick;
blanchet
parents: 35284
diff changeset
   319
    | Normal ([], _, _) => "none"
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   320
    | Normal _ => "genuine"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   321
    | TimedOut _ => "unknown"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   322
    | Interrupted _ => "unknown"
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   323
    | Error (s, _) => error ("Kodkod error: " ^ s)
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   324
  end
33980
a28733ef3a82 export symbols from Minipick (so I can use them in other programs)
blanchet
parents: 33232
diff changeset
   325
33192
08a39a957ed7 added Nitpick's theory and ML files to Isabelle/HOL;
blanchet
parents:
diff changeset
   326
end;