| author | wenzelm | 
| Mon, 03 Feb 2025 19:53:13 +0100 | |
| changeset 82072 | 4df95adf6b66 | 
| parent 74844 | 90242c744a1a | 
| permissions | -rw-r--r-- | 
| 45035 | 1 | (* Title: HOL/Nitpick_Examples/minipick.ML | 
| 2 | Author: Jasmin Blanchette, TU Muenchen | |
| 3 | Copyright 2009-2010 | |
| 4 | ||
| 5 | Finite model generation for HOL formulas using Kodkod, minimalistic version. | |
| 6 | *) | |
| 7 | ||
| 8 | signature MINIPICK = | |
| 9 | sig | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 10 | val minipick : Proof.context -> int -> term -> string | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 11 | val minipick_expect : Proof.context -> string -> int -> term -> unit | 
| 45035 | 12 | end; | 
| 13 | ||
| 14 | structure Minipick : MINIPICK = | |
| 15 | struct | |
| 16 | ||
| 17 | open Kodkod | |
| 18 | open Nitpick_Util | |
| 19 | open Nitpick_HOL | |
| 20 | open Nitpick_Peephole | |
| 21 | open Nitpick_Kodkod | |
| 22 | ||
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 23 | datatype rep = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 24 | S_Rep | | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 25 | R_Rep of bool | 
| 45035 | 26 | |
| 74399 | 27 | fun check_type ctxt raw_infinite \<^Type>\<open>fun T1 T2\<close> = | 
| 28 | List.app (check_type ctxt raw_infinite) [T1, T2] | |
| 29 | | check_type ctxt raw_infinite \<^Type>\<open>prod T1 T2\<close> = | |
| 30 | List.app (check_type ctxt raw_infinite) [T1, T2] | |
| 31 | | check_type _ _ \<^Type>\<open>bool\<close> = () | |
| 69597 | 32 |   | check_type _ _ (TFree (_, \<^sort>\<open>{}\<close>)) = ()
 | 
| 33 | | check_type _ _ (TFree (_, \<^sort>\<open>HOL.type\<close>)) = () | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 34 | | check_type ctxt raw_infinite T = | 
| 46092 | 35 | if raw_infinite T then | 
| 36 | () | |
| 37 | else | |
| 38 |       error ("Not supported: Type " ^ quote (Syntax.string_of_typ ctxt T) ^ ".")
 | |
| 45035 | 39 | |
| 74399 | 40 | fun atom_schema_of S_Rep card \<^Type>\<open>fun T1 T2\<close> = | 
| 45035 | 41 | replicate_list (card T1) (atom_schema_of S_Rep card T2) | 
| 74399 | 42 | | atom_schema_of (R_Rep true) card \<^Type>\<open>fun T1 \<^Type>\<open>bool\<close>\<close> = | 
| 45035 | 43 | atom_schema_of S_Rep card T1 | 
| 74399 | 44 | | atom_schema_of (rep as R_Rep _) card \<^Type>\<open>fun T1 T2\<close> = | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 45 | atom_schema_of S_Rep card T1 @ atom_schema_of rep card T2 | 
| 74399 | 46 | | atom_schema_of _ card \<^Type>\<open>prod T1 T2\<close> = | 
| 47 | maps (atom_schema_of S_Rep card) [T1, T2] | |
| 45035 | 48 | | atom_schema_of _ card T = [card T] | 
| 49 | val arity_of = length ooo atom_schema_of | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 50 | val atom_seqs_of = map (AtomSeq o rpair 0) ooo atom_schema_of | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 51 | val atom_seq_product_of = foldl1 Product ooo atom_seqs_of | 
| 45035 | 52 | |
| 53 | fun index_for_bound_var _ [_] 0 = 0 | |
| 54 | | index_for_bound_var card (_ :: Ts) 0 = | |
| 55 | index_for_bound_var card Ts 0 + arity_of S_Rep card (hd Ts) | |
| 56 | | index_for_bound_var card Ts n = index_for_bound_var card (tl Ts) (n - 1) | |
| 57 | fun vars_for_bound_var card R Ts j = | |
| 58 | map (curry Var 1) (index_seq (index_for_bound_var card Ts j) | |
| 59 | (arity_of R card (nth Ts j))) | |
| 60 | val rel_expr_for_bound_var = foldl1 Product oooo vars_for_bound_var | |
| 61 | fun decls_for R card Ts T = | |
| 62 | map2 (curry DeclOne o pair 1) | |
| 63 | (index_seq (index_for_bound_var card (T :: Ts) 0) | |
| 64 | (arity_of R card (nth (T :: Ts) 0))) | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 65 | (atom_seqs_of R card T) | 
| 45035 | 66 | |
| 67 | val atom_product = foldl1 Product o map Atom | |
| 68 | ||
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 69 | val false_atom_num = 0 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 70 | val true_atom_num = 1 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 71 | val false_atom = Atom false_atom_num | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 72 | val true_atom = Atom true_atom_num | 
| 45035 | 73 | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 74 | fun kodkod_formula_from_term ctxt total card complete concrete frees = | 
| 45035 | 75 | let | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 76 | fun F_from_S_rep (SOME false) r = Not (RelEq (r, false_atom)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 77 | | F_from_S_rep _ r = RelEq (r, true_atom) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 78 | fun S_rep_from_F NONE f = RelIf (f, true_atom, false_atom) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 79 | | S_rep_from_F (SOME true) f = RelIf (f, true_atom, None) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 80 | | S_rep_from_F (SOME false) f = RelIf (Not f, false_atom, None) | 
| 74399 | 81 | fun R_rep_from_S_rep \<^Type>\<open>fun T1 T2\<close> r = | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 82 | if total andalso T2 = bool_T then | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 83 | let | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 84 | val jss = atom_schema_of S_Rep card T1 |> map (rpair 0) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 85 | |> all_combinations | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 86 | in | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 87 | map2 (fn i => fn js => | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 88 | (* | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 89 | RelIf (F_from_S_rep NONE (Project (r, [Num i])), | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 90 | atom_product js, empty_n_ary_rel (length js)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 91 | *) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 92 | Join (Project (r, [Num i]), | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 93 | atom_product (false_atom_num :: js)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 94 | ) (index_seq 0 (length jss)) jss | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 95 | |> foldl1 Union | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 96 | end | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 97 | else | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 98 | let | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 99 | val jss = atom_schema_of S_Rep card T1 |> map (rpair 0) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 100 | |> all_combinations | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 101 | val arity2 = arity_of S_Rep card T2 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 102 | in | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 103 | map2 (fn i => fn js => | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 104 | Product (atom_product js, | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 105 | Project (r, num_seq (i * arity2) arity2) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 106 | |> R_rep_from_S_rep T2)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 107 | (index_seq 0 (length jss)) jss | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 108 | |> foldl1 Union | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 109 | end | 
| 45035 | 110 | | R_rep_from_S_rep _ r = r | 
| 74399 | 111 | fun S_rep_from_R_rep Ts (T as \<^Type>\<open>fun _ _\<close>) r = | 
| 45035 | 112 | Comprehension (decls_for S_Rep card Ts T, | 
| 113 | RelEq (R_rep_from_S_rep T | |
| 114 | (rel_expr_for_bound_var card S_Rep (T :: Ts) 0), r)) | |
| 115 | | S_rep_from_R_rep _ _ r = r | |
| 74399 | 116 | fun partial_eq pos Ts \<^Type>\<open>fun T1 T2\<close> t1 t2 = | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 117 |         HOLogic.mk_all ("x", T1,
 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 118 | HOLogic.eq_const T2 $ (incr_boundvars 1 t1 $ Bound 0) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 119 | $ (incr_boundvars 1 t2 $ Bound 0)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 120 | |> to_F (SOME pos) Ts | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 121 | | partial_eq pos Ts T t1 t2 = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 122 | if pos andalso not (concrete T) then | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 123 | False | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 124 | else | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
55208diff
changeset | 125 | (t1, t2) |> apply2 (to_R_rep Ts) | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 126 | |> (if pos then Some o Intersect else Lone o Union) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 127 | and to_F pos Ts t = | 
| 45035 | 128 | (case t of | 
| 74399 | 129 | \<^Const_>\<open>Not for t1\<close> => Not (to_F (Option.map not pos) Ts t1) | 
| 130 | | \<^Const_>\<open>False\<close> => False | |
| 131 | | \<^Const_>\<open>True\<close> => True | |
| 132 | | \<^Const_>\<open>All _ for \<open>Abs (_, T, t')\<close>\<close> => | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 133 | if pos = SOME true andalso not (complete T) then False | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 134 | else All (decls_for S_Rep card Ts T, to_F pos (T :: Ts) t') | 
| 74399 | 135 | | (t0 as \<^Const_>\<open>All _\<close>) $ t1 => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 136 | to_F pos Ts (t0 $ eta_expand Ts t1 1) | 
| 74399 | 137 | | \<^Const_>\<open>Ex _ for \<open>Abs (_, T, t')\<close>\<close> => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 138 | if pos = SOME false andalso not (complete T) then True | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 139 | else Exist (decls_for S_Rep card Ts T, to_F pos (T :: Ts) t') | 
| 74399 | 140 | | (t0 as \<^Const_>\<open>Ex _\<close>) $ t1 => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 141 | to_F pos Ts (t0 $ eta_expand Ts t1 1) | 
| 74399 | 142 | | \<^Const_>\<open>HOL.eq T for t1 t2\<close> => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 143 | (case pos of | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 144 | NONE => RelEq (to_R_rep Ts t1, to_R_rep Ts t2) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 145 | | SOME pos => partial_eq pos Ts T t1 t2) | 
| 74399 | 146 | | \<^Const_>\<open>less_eq \<^Type>\<open>fun T' \<^Type>\<open>bool\<close>\<close> for t1 t2\<close> => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 147 | (case pos of | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 148 | NONE => Subset (to_R_rep Ts t1, to_R_rep Ts t2) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 149 | | SOME true => | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 150 | Subset (Difference (atom_seq_product_of S_Rep card T', | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 151 | Join (to_R_rep Ts t1, false_atom)), | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 152 | Join (to_R_rep Ts t2, true_atom)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 153 | | SOME false => | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 154 | Subset (Join (to_R_rep Ts t1, true_atom), | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 155 | Difference (atom_seq_product_of S_Rep card T', | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 156 | Join (to_R_rep Ts t2, false_atom)))) | 
| 74399 | 157 | | \<^Const_>\<open>conj for t1 t2\<close> => And (to_F pos Ts t1, to_F pos Ts t2) | 
| 158 | | \<^Const_>\<open>disj for t1 t2\<close> => Or (to_F pos Ts t1, to_F pos Ts t2) | |
| 159 | | \<^Const_>\<open>implies for t1 t2\<close> => | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 160 | Implies (to_F (Option.map not pos) Ts t1, to_F pos Ts t2) | 
| 74399 | 161 | | \<^Const_>\<open>Set.member _ for t1 t2\<close> => to_F pos Ts (t2 $ t1) | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 162 | | t1 $ t2 => | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 163 | (case pos of | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 164 | NONE => Subset (to_S_rep Ts t2, to_R_rep Ts t1) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 165 | | SOME pos => | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 166 | let | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 167 | val kt1 = to_R_rep Ts t1 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 168 | val kt2 = to_S_rep Ts t2 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 169 | val kT = atom_seq_product_of S_Rep card (fastype_of1 (Ts, t2)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 170 | in | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 171 | if pos then | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 172 | Not (Subset (kt2, Difference (kT, Join (kt1, true_atom)))) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 173 | else | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 174 | Subset (kt2, Difference (kT, Join (kt1, false_atom))) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 175 | end) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 176 | | _ => raise SAME ()) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 177 | handle SAME () => F_from_S_rep pos (to_R_rep Ts t) | 
| 45035 | 178 | and to_S_rep Ts t = | 
| 179 | case t of | |
| 74399 | 180 | \<^Const_>\<open>Pair _ _ for t1 t2\<close> => Product (to_S_rep Ts t1, to_S_rep Ts t2) | 
| 181 | | \<^Const_>\<open>Pair _ _ for _\<close> => to_S_rep Ts (eta_expand Ts t 1) | |
| 182 | | \<^Const_>\<open>Pair _ _\<close> => to_S_rep Ts (eta_expand Ts t 2) | |
| 183 | | \<^Const_>\<open>fst _ _ for t1\<close> => | |
| 45035 | 184 | let val fst_arity = arity_of S_Rep card (fastype_of1 (Ts, t)) in | 
| 185 | Project (to_S_rep Ts t1, num_seq 0 fst_arity) | |
| 186 | end | |
| 74399 | 187 | | \<^Const_>\<open>fst _ _\<close> => to_S_rep Ts (eta_expand Ts t 1) | 
| 188 | | \<^Const_>\<open>snd _ _ for t1\<close> => | |
| 45035 | 189 | let | 
| 190 | val pair_arity = arity_of S_Rep card (fastype_of1 (Ts, t1)) | |
| 191 | val snd_arity = arity_of S_Rep card (fastype_of1 (Ts, t)) | |
| 192 | val fst_arity = pair_arity - snd_arity | |
| 193 | in Project (to_S_rep Ts t1, num_seq fst_arity snd_arity) end | |
| 74399 | 194 | | \<^Const_>\<open>snd _ _\<close> => to_S_rep Ts (eta_expand Ts t 1) | 
| 45035 | 195 | | Bound j => rel_expr_for_bound_var card S_Rep Ts j | 
| 196 | | _ => S_rep_from_R_rep Ts (fastype_of1 (Ts, t)) (to_R_rep Ts t) | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 197 | and partial_set_op swap1 swap2 op1 op2 Ts t1 t2 = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 198 | let | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 199 | val kt1 = to_R_rep Ts t1 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 200 | val kt2 = to_R_rep Ts t2 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 201 | val (a11, a21) = (false_atom, true_atom) |> swap1 ? swap | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 202 | val (a12, a22) = (false_atom, true_atom) |> swap2 ? swap | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 203 | in | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 204 | Union (Product (op1 (Join (kt1, a11), Join (kt2, a12)), true_atom), | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 205 | Product (op2 (Join (kt1, a21), Join (kt2, a22)), false_atom)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 206 | end | 
| 45035 | 207 | and to_R_rep Ts t = | 
| 208 | (case t of | |
| 74399 | 209 | \<^Const_>\<open>Not\<close> => to_R_rep Ts (eta_expand Ts t 1) | 
| 210 | | \<^Const_>\<open>All _\<close> => to_R_rep Ts (eta_expand Ts t 1) | |
| 211 | | \<^Const_>\<open>Ex _\<close> => to_R_rep Ts (eta_expand Ts t 1) | |
| 212 | | \<^Const_>\<open>HOL.eq _ for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | |
| 213 | | \<^Const_>\<open>HOL.eq _\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 214 | | \<^Const_>\<open>less_eq \<^Type>\<open>fun _ \<^Type>\<open>bool\<close>\<close> for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | |
| 215 | | \<^Const_>\<open>less_eq _\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 216 | | \<^Const_>\<open>conj for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | |
| 217 | | \<^Const_>\<open>conj\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 218 | | \<^Const_>\<open>disj for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | |
| 219 | | \<^Const_>\<open>disj\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 220 | | \<^Const_>\<open>implies for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | |
| 221 | | \<^Const_>\<open>implies\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 222 | | \<^Const_>\<open>Set.member _ for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | |
| 223 | | \<^Const_>\<open>Set.member _\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 224 | | \<^Const_>\<open>Collect _ for t'\<close> => to_R_rep Ts t' | |
| 225 | | \<^Const_>\<open>Collect _\<close> => to_R_rep Ts (eta_expand Ts t 1) | |
| 226 | | \<^Const_>\<open>bot \<open>T as \<^Type>\<open>fun T' \<^Type>\<open>bool\<close>\<close>\<close>\<close> => | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 227 | if total then empty_n_ary_rel (arity_of (R_Rep total) card T) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 228 | else Product (atom_seq_product_of (R_Rep total) card T', false_atom) | 
| 74399 | 229 | | \<^Const_>\<open>top \<open>T as \<^Type>\<open>fun T' \<^Type>\<open>bool\<close>\<close>\<close>\<close> => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 230 | if total then atom_seq_product_of (R_Rep total) card T | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 231 | else Product (atom_seq_product_of (R_Rep total) card T', true_atom) | 
| 74399 | 232 | | \<^Const_>\<open>insert T for t1 t2\<close> => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 233 | if total then | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 234 | Union (to_S_rep Ts t1, to_R_rep Ts t2) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 235 | else | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 236 | let | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 237 | val kt1 = to_S_rep Ts t1 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 238 | val kt2 = to_R_rep Ts t2 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 239 | in | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 240 | RelIf (Some kt1, | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 241 | if arity_of S_Rep card T = 1 then | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 242 | Override (kt2, Product (kt1, true_atom)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 243 | else | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 244 | Union (Difference (kt2, Product (kt1, false_atom)), | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 245 | Product (kt1, true_atom)), | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 246 | Difference (kt2, Product (atom_seq_product_of S_Rep card T, | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 247 | false_atom))) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 248 | end | 
| 74399 | 249 | | \<^Const_>\<open>insert _ for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | 
| 250 | | \<^Const_>\<open>insert _\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 251 | | Const (\<^const_name>\<open>trancl\<close>, (* FIXME proper type!? *) | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 252 | Type (_, [Type (_, [Type (_, [T', _]), _]), _])) $ t1 => | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 253 | if arity_of S_Rep card T' = 1 then | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 254 | if total then | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 255 | Closure (to_R_rep Ts t1) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 256 | else | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 257 | let | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 258 | val kt1 = to_R_rep Ts t1 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 259 | val true_core_kt = Closure (Join (kt1, true_atom)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 260 | val kTx = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 261 | atom_seq_product_of S_Rep card (HOLogic.mk_prodT (`I T')) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 262 | val false_mantle_kt = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 263 | Difference (kTx, | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 264 | Closure (Difference (kTx, Join (kt1, false_atom)))) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 265 | in | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 266 | Union (Product (Difference (false_mantle_kt, true_core_kt), | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 267 | false_atom), | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 268 | Product (true_core_kt, true_atom)) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 269 | end | 
| 45035 | 270 | else | 
| 46092 | 271 | error "Not supported: Transitive closure for function or pair type." | 
| 74399 | 272 | | \<^Const_>\<open>trancl _\<close> => to_R_rep Ts (eta_expand Ts t 1) | 
| 273 | | \<^Const_>\<open>inf \<^Type>\<open>fun _ \<^Type>\<open>bool\<close>\<close> for t1 t2\<close> => | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 274 | if total then Intersect (to_R_rep Ts t1, to_R_rep Ts t2) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 275 | else partial_set_op true true Intersect Union Ts t1 t2 | 
| 74399 | 276 | | \<^Const_>\<open>inf _ for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | 
| 277 | | \<^Const_>\<open>inf _\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 278 | | \<^Const_>\<open>sup \<^Type>\<open>fun _ \<^Type>\<open>bool\<close>\<close> for t1 t2\<close> => | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 279 | if total then Union (to_R_rep Ts t1, to_R_rep Ts t2) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 280 | else partial_set_op true true Union Intersect Ts t1 t2 | 
| 74399 | 281 | | \<^Const_>\<open>sup _ for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | 
| 282 | | \<^Const_>\<open>sup _\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 283 | | \<^Const_>\<open>minus \<^Type>\<open>fun _ \<^Type>\<open>bool\<close>\<close> for t1 t2\<close> => | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 284 | if total then Difference (to_R_rep Ts t1, to_R_rep Ts t2) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 285 | else partial_set_op true false Intersect Union Ts t1 t2 | 
| 74399 | 286 | | \<^Const_>\<open>minus \<^Type>\<open>fun _ \<^Type>\<open>bool\<close>\<close> for _\<close> => to_R_rep Ts (eta_expand Ts t 1) | 
| 287 | | \<^Const_>\<open>minus \<^Type>\<open>fun _ \<^Type>\<open>bool\<close>\<close>\<close> => to_R_rep Ts (eta_expand Ts t 2) | |
| 288 | | \<^Const_>\<open>Pair _ _ for _ _\<close> => to_S_rep Ts t | |
| 289 | | \<^Const_>\<open>Pair _ _ for _\<close> => to_S_rep Ts t | |
| 290 | | \<^Const_>\<open>Pair _ _\<close> => to_S_rep Ts t | |
| 291 | | \<^Const_>\<open>fst _ _ for _\<close> => raise SAME () | |
| 292 | | \<^Const_>\<open>fst _ _\<close> => raise SAME () | |
| 293 | | \<^Const_>\<open>snd _ _ for _\<close> => raise SAME () | |
| 294 | | \<^Const_>\<open>snd _ _\<close> => raise SAME () | |
| 295 | | \<^Const_>\<open>False\<close> => false_atom | |
| 296 | | \<^Const_>\<open>True\<close> => true_atom | |
| 45035 | 297 | | Free (x as (_, T)) => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 298 | Rel (arity_of (R_Rep total) card T, find_index (curry (op =) x) frees) | 
| 46092 | 299 | | Term.Var _ => error "Not supported: Schematic variables." | 
| 45035 | 300 | | Bound _ => raise SAME () | 
| 301 | | Abs (_, T, t') => | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 302 | (case (total, fastype_of1 (T :: Ts, t')) of | 
| 74399 | 303 | (true, \<^Type>\<open>bool\<close>) => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 304 | Comprehension (decls_for S_Rep card Ts T, to_F NONE (T :: Ts) t') | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 305 | | (_, T') => | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 306 | Comprehension (decls_for S_Rep card Ts T @ | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 307 | decls_for (R_Rep total) card (T :: Ts) T', | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 308 | Subset (rel_expr_for_bound_var card (R_Rep total) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 309 | (T' :: T :: Ts) 0, | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 310 | to_R_rep (T :: Ts) t'))) | 
| 45035 | 311 | | t1 $ t2 => | 
| 312 | (case fastype_of1 (Ts, t) of | |
| 74399 | 313 | \<^Type>\<open>bool\<close> => | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 314 | if total then | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 315 | S_rep_from_F NONE (to_F NONE Ts t) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 316 | else | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 317 | RelIf (to_F (SOME true) Ts t, true_atom, | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 318 | RelIf (Not (to_F (SOME false) Ts t), false_atom, | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 319 | None)) | 
| 45035 | 320 | | T => | 
| 321 | let val T2 = fastype_of1 (Ts, t2) in | |
| 322 | case arity_of S_Rep card T2 of | |
| 323 | 1 => Join (to_S_rep Ts t2, to_R_rep Ts t1) | |
| 324 | | arity2 => | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 325 | let val res_arity = arity_of (R_Rep total) card T in | 
| 45035 | 326 | Project (Intersect | 
| 327 | (Product (to_S_rep Ts t2, | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 328 | atom_seq_product_of (R_Rep total) card T), | 
| 45035 | 329 | to_R_rep Ts t1), | 
| 330 | num_seq arity2 res_arity) | |
| 331 | end | |
| 332 | end) | |
| 46092 | 333 |        | _ => error ("Not supported: Term " ^
 | 
| 334 | quote (Syntax.string_of_term ctxt t) ^ ".")) | |
| 45035 | 335 | handle SAME () => R_rep_from_S_rep (fastype_of1 (Ts, t)) (to_S_rep Ts t) | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 336 | in to_F (if total then NONE else SOME true) [] end | 
| 45035 | 337 | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 338 | fun bound_for_free total card i (s, T) = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 339 | let val js = atom_schema_of (R_Rep total) card T in | 
| 45035 | 340 | ([((length js, i), s)], | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 341 | [TupleSet [], atom_schema_of (R_Rep total) card T |> map (rpair 0) | 
| 45035 | 342 | |> tuple_set_from_atom_schema]) | 
| 343 | end | |
| 344 | ||
| 74399 | 345 | fun declarative_axiom_for_rel_expr total card Ts \<^Type>\<open>fun T1 T2\<close> r = | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 346 | if total andalso body_type T2 = bool_T then | 
| 45035 | 347 | True | 
| 348 | else | |
| 349 | All (decls_for S_Rep card Ts T1, | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 350 | declarative_axiom_for_rel_expr total card (T1 :: Ts) T2 | 
| 45035 | 351 | (List.foldl Join r (vars_for_bound_var card S_Rep (T1 :: Ts) 0))) | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 352 | | declarative_axiom_for_rel_expr total _ _ _ r = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 353 | (if total then One else Lone) r | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 354 | fun declarative_axiom_for_free total card i (_, T) = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 355 | declarative_axiom_for_rel_expr total card [] T | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 356 | (Rel (arity_of (R_Rep total) card T, i)) | 
| 45035 | 357 | |
| 46092 | 358 | (* Hack to make the old code work as is with sets. *) | 
| 74399 | 359 | fun unsetify_type \<^Type>\<open>set T\<close> = unsetify_type T --> bool_T | 
| 46092 | 360 | | unsetify_type (Type (s, Ts)) = Type (s, map unsetify_type Ts) | 
| 361 | | unsetify_type T = T | |
| 362 | ||
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 363 | fun kodkod_problem_from_term ctxt total raw_card raw_infinite t = | 
| 45035 | 364 | let | 
| 45128 
5af3a3203a76
discontinued obsolete alias structure ProofContext;
 wenzelm parents: 
45062diff
changeset | 365 | val thy = Proof_Context.theory_of ctxt | 
| 74399 | 366 | fun card \<^Type>\<open>fun T1 T2\<close> = reasonable_power (card T2) (card T1) | 
| 367 | | card \<^Type>\<open>prod T1 T2\<close> = card T1 * card T2 | |
| 368 | | card \<^Type>\<open>bool\<close> = 2 | |
| 45035 | 369 | | card T = Int.max (1, raw_card T) | 
| 74399 | 370 | fun complete \<^Type>\<open>fun T1 T2\<close> = concrete T1 andalso complete T2 | 
| 371 | | complete \<^Type>\<open>prod T1 T2\<close> = complete T1 andalso complete T2 | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 372 | | complete T = not (raw_infinite T) | 
| 74399 | 373 | and concrete \<^Type>\<open>fun T1 T2\<close> = complete T1 andalso concrete T2 | 
| 374 | | concrete \<^Type>\<open>prod T1 T2\<close> = concrete T1 andalso concrete T2 | |
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 375 | | concrete _ = true | 
| 46092 | 376 | val neg_t = | 
| 74399 | 377 | \<^Const>\<open>Not\<close> $ Object_Logic.atomize_term ctxt t | 
| 46092 | 378 | |> map_types unsetify_type | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 379 | val _ = fold_types (K o check_type ctxt raw_infinite) neg_t () | 
| 45035 | 380 | val frees = Term.add_frees neg_t [] | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 381 | val bounds = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 382 | map2 (bound_for_free total card) (index_seq 0 (length frees)) frees | 
| 45035 | 383 | val declarative_axioms = | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 384 | map2 (declarative_axiom_for_free total card) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 385 | (index_seq 0 (length frees)) frees | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 386 | val formula = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 387 | neg_t |> kodkod_formula_from_term ctxt total card complete concrete frees | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 388 | |> fold_rev (curry And) declarative_axioms | 
| 45035 | 389 | val univ_card = univ_card 0 0 0 bounds formula | 
| 390 | in | |
| 391 |     {comment = "", settings = [], univ_card = univ_card, tuple_assigns = [],
 | |
| 392 | bounds = bounds, int_bounds = [], expr_assigns = [], formula = formula} | |
| 393 | end | |
| 394 | ||
| 395 | fun solve_any_kodkod_problem thy problems = | |
| 396 | let | |
| 55199 | 397 |     val {debug, overlord, timeout, ...} = Nitpick_Commands.default_params thy []
 | 
| 74844 
90242c744a1a
maintain option kodkod_scala within theory context, to allow local modification;
 wenzelm parents: 
74399diff
changeset | 398 | val kodkod_scala = Config.get_global thy Kodkod.kodkod_scala | 
| 73387 | 399 | val deadline = Timeout.end_time timeout | 
| 45035 | 400 | val max_threads = 1 | 
| 401 | val max_solutions = 1 | |
| 402 | in | |
| 74844 
90242c744a1a
maintain option kodkod_scala within theory context, to allow local modification;
 wenzelm parents: 
74399diff
changeset | 403 | case solve_any_problem kodkod_scala debug overlord deadline max_threads max_solutions | 
| 45035 | 404 | problems of | 
| 72329 
6255e532aa36
obsolete --- Java is always present via component;
 wenzelm parents: 
72328diff
changeset | 405 | Normal ([], _, _) => "none" | 
| 45035 | 406 | | Normal _ => "genuine" | 
| 407 | | TimedOut _ => "unknown" | |
| 408 |     | Error (s, _) => error ("Kodkod error: " ^ s)
 | |
| 409 | end | |
| 410 | ||
| 74399 | 411 | val default_raw_infinite = member (op =) [\<^Type>\<open>nat\<close>, \<^Type>\<open>int\<close>] | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 412 | |
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 413 | fun minipick ctxt n t = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 414 | let | 
| 45128 
5af3a3203a76
discontinued obsolete alias structure ProofContext;
 wenzelm parents: 
45062diff
changeset | 415 | val thy = Proof_Context.theory_of ctxt | 
| 55199 | 416 |     val {total_consts, ...} = Nitpick_Commands.default_params thy []
 | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 417 | val totals = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 418 | total_consts |> Option.map single |> the_default [true, false] | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 419 | fun problem_for (total, k) = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 420 | kodkod_problem_from_term ctxt total (K k) default_raw_infinite t | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 421 | in | 
| 55208 | 422 | (totals, 1 upto n) | 
| 45062 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 423 | |-> map_product pair | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 424 | |> map problem_for | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 425 | |> solve_any_kodkod_problem (Proof_Context.theory_of ctxt) | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 426 | end | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 427 | |
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 428 | fun minipick_expect ctxt expect n t = | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 429 | if getenv "KODKODI" <> "" then | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 430 | if minipick ctxt n t = expect then () | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 431 |     else error ("\"minipick_expect\" expected " ^ quote expect)
 | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 432 | else | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 433 | () | 
| 
9598cada31b3
first step towards extending Minipick with more translations
 blanchet parents: 
45035diff
changeset | 434 | |
| 45035 | 435 | end; |