author | blanchet |
Tue, 24 Nov 2009 18:36:18 +0100 | |
changeset 33892 | 3937da7e13ea |
parent 33232 | f93390060bbe |
child 33980 | a28733ef3a82 |
permissions | -rw-r--r-- |
33192 | 1 |
(* Title: HOL/Nitpick/Tools/minipick.ML |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
3 |
Copyright 2009 |
|
4 |
||
5 |
Finite model generation for HOL formulas using Kodkod, minimalistic version. |
|
6 |
*) |
|
7 |
||
8 |
signature MINIPICK = |
|
9 |
sig |
|
10 |
val pick_nits_in_term : Proof.context -> (typ -> int) -> term -> string |
|
11 |
end; |
|
12 |
||
13 |
structure Minipick : MINIPICK = |
|
14 |
struct |
|
15 |
||
16 |
open Kodkod |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
17 |
open Nitpick_Util |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
18 |
open Nitpick_HOL |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
19 |
open Nitpick_Peephole |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
20 |
open Nitpick_Kodkod |
33192 | 21 |
|
22 |
(* theory -> typ -> unit *) |
|
23 |
fun check_type ctxt (Type ("fun", Ts)) = List.app (check_type ctxt) Ts |
|
24 |
| check_type ctxt (Type ("*", Ts)) = List.app (check_type ctxt) Ts |
|
25 |
| check_type _ @{typ bool} = () |
|
26 |
| check_type _ (TFree (_, @{sort "{}"})) = () |
|
27 |
| check_type _ (TFree (_, @{sort HOL.type})) = () |
|
28 |
| check_type ctxt T = |
|
29 |
raise NOT_SUPPORTED ("type " ^ quote (Syntax.string_of_typ ctxt T)) |
|
30 |
||
31 |
(* (typ -> int) -> typ -> int *) |
|
32 |
fun atom_schema_of_one scope (Type ("fun", [T1, T2])) = |
|
33 |
replicate_list (scope T1) (atom_schema_of_one scope T2) |
|
34 |
| atom_schema_of_one scope (Type ("*", [T1, T2])) = |
|
35 |
atom_schema_of_one scope T1 @ atom_schema_of_one scope T2 |
|
36 |
| atom_schema_of_one scope T = [scope T] |
|
37 |
fun atom_schema_of_set scope (Type ("fun", [T1, @{typ bool}])) = |
|
38 |
atom_schema_of_one scope T1 |
|
39 |
| atom_schema_of_set scope (Type ("fun", [T1, T2])) = |
|
40 |
atom_schema_of_one scope T1 @ atom_schema_of_set scope T2 |
|
41 |
| atom_schema_of_set scope T = atom_schema_of_one scope T |
|
42 |
val arity_of_one = length oo atom_schema_of_one |
|
43 |
val arity_of_set = length oo atom_schema_of_set |
|
44 |
||
45 |
(* (typ -> int) -> typ list -> int -> int *) |
|
46 |
fun index_for_bound_var _ [_] 0 = 0 |
|
47 |
| index_for_bound_var scope (_ :: Ts) 0 = |
|
48 |
index_for_bound_var scope Ts 0 + arity_of_one scope (hd Ts) |
|
49 |
| index_for_bound_var scope Ts n = index_for_bound_var scope (tl Ts) (n - 1) |
|
50 |
(* (typ -> int) -> typ list -> int -> rel_expr list *) |
|
51 |
fun one_vars_for_bound_var scope Ts j = |
|
52 |
map (curry Var 1) (index_seq (index_for_bound_var scope Ts j) |
|
53 |
(arity_of_one scope (nth Ts j))) |
|
54 |
fun set_vars_for_bound_var scope Ts j = |
|
55 |
map (curry Var 1) (index_seq (index_for_bound_var scope Ts j) |
|
56 |
(arity_of_set scope (nth Ts j))) |
|
57 |
(* (typ -> int) -> typ list -> typ -> decl list *) |
|
58 |
fun decls_for_one scope Ts T = |
|
59 |
map2 (curry DeclOne o pair 1) |
|
60 |
(index_seq (index_for_bound_var scope (T :: Ts) 0) |
|
61 |
(arity_of_one scope (nth (T :: Ts) 0))) |
|
62 |
(map (AtomSeq o rpair 0) (atom_schema_of_one scope T)) |
|
63 |
fun decls_for_set scope Ts T = |
|
64 |
map2 (curry DeclOne o pair 1) |
|
65 |
(index_seq (index_for_bound_var scope (T :: Ts) 0) |
|
66 |
(arity_of_set scope (nth (T :: Ts) 0))) |
|
67 |
(map (AtomSeq o rpair 0) (atom_schema_of_set scope T)) |
|
68 |
||
69 |
(* int list -> rel_expr *) |
|
70 |
val atom_product = foldl1 Product o map Atom |
|
71 |
||
72 |
val false_atom = Atom 0 |
|
73 |
val true_atom = Atom 1 |
|
74 |
||
75 |
(* rel_expr -> formula *) |
|
76 |
fun formula_from_atom r = RelEq (r, true_atom) |
|
77 |
(* formula -> rel_expr *) |
|
78 |
fun atom_from_formula f = RelIf (f, true_atom, false_atom) |
|
79 |
||
80 |
(* Proof.context -> (typ -> int) -> styp list -> term -> formula *) |
|
81 |
fun kodkod_formula_for_term ctxt scope frees = |
|
82 |
let |
|
83 |
(* typ list -> int -> rel_expr *) |
|
84 |
val one_from_bound_var = foldl1 Product oo one_vars_for_bound_var scope |
|
85 |
val set_from_bound_var = foldl1 Product oo set_vars_for_bound_var scope |
|
86 |
(* typ -> rel_expr -> rel_expr *) |
|
87 |
fun set_from_one (T as Type ("fun", [T1, @{typ bool}])) r = |
|
88 |
let |
|
89 |
val jss = atom_schema_of_one scope T1 |> map (rpair 0) |
|
90 |
|> all_combinations |
|
91 |
in |
|
92 |
map2 (fn i => fn js => |
|
93 |
RelIf (RelEq (Project (r, [Num i]), true_atom), |
|
94 |
atom_product js, empty_n_ary_rel (length js))) |
|
95 |
(index_seq 0 (length jss)) jss |
|
96 |
|> foldl1 Union |
|
97 |
end |
|
98 |
| set_from_one (Type ("fun", [T1, T2])) r = |
|
99 |
let |
|
100 |
val jss = atom_schema_of_one scope T1 |> map (rpair 0) |
|
101 |
|> all_combinations |
|
102 |
val arity2 = arity_of_one scope T2 |
|
103 |
in |
|
104 |
map2 (fn i => fn js => |
|
105 |
Product (atom_product js, |
|
106 |
Project (r, num_seq (i * arity2) arity2) |
|
107 |
|> set_from_one T2)) |
|
108 |
(index_seq 0 (length jss)) jss |
|
109 |
|> foldl1 Union |
|
110 |
end |
|
111 |
| set_from_one _ r = r |
|
112 |
(* typ list -> typ -> rel_expr -> rel_expr *) |
|
113 |
fun one_from_set Ts (T as Type ("fun", _)) r = |
|
114 |
Comprehension (decls_for_one scope Ts T, |
|
115 |
RelEq (set_from_one T (one_from_bound_var (T :: Ts) 0), |
|
116 |
r)) |
|
117 |
| one_from_set _ _ r = r |
|
118 |
(* typ list -> term -> formula *) |
|
119 |
fun to_f Ts t = |
|
120 |
(case t of |
|
121 |
@{const Not} $ t1 => Not (to_f Ts t1) |
|
122 |
| @{const False} => False |
|
123 |
| @{const True} => True |
|
124 |
| Const (@{const_name All}, _) $ Abs (s, T, t') => |
|
125 |
All (decls_for_one scope Ts T, to_f (T :: Ts) t') |
|
126 |
| (t0 as Const (@{const_name All}, _)) $ t1 => |
|
127 |
to_f Ts (t0 $ eta_expand Ts t1 1) |
|
128 |
| Const (@{const_name Ex}, _) $ Abs (s, T, t') => |
|
129 |
Exist (decls_for_one scope Ts T, to_f (T :: Ts) t') |
|
130 |
| (t0 as Const (@{const_name Ex}, _)) $ t1 => |
|
131 |
to_f Ts (t0 $ eta_expand Ts t1 1) |
|
132 |
| Const (@{const_name "op ="}, _) $ t1 $ t2 => |
|
133 |
RelEq (to_set Ts t1, to_set Ts t2) |
|
134 |
| Const (@{const_name ord_class.less_eq}, |
|
135 |
Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ t1 $ t2 => |
|
136 |
Subset (to_set Ts t1, to_set Ts t2) |
|
137 |
| @{const "op &"} $ t1 $ t2 => And (to_f Ts t1, to_f Ts t2) |
|
138 |
| @{const "op |"} $ t1 $ t2 => Or (to_f Ts t1, to_f Ts t2) |
|
139 |
| @{const "op -->"} $ t1 $ t2 => Implies (to_f Ts t1, to_f Ts t2) |
|
140 |
| t1 $ t2 => Subset (to_one Ts t2, to_set Ts t1) |
|
141 |
| Free _ => raise SAME () |
|
142 |
| Term.Var _ => raise SAME () |
|
143 |
| Bound _ => raise SAME () |
|
144 |
| Const (s, _) => raise NOT_SUPPORTED ("constant " ^ quote s) |
|
145 |
| _ => raise TERM ("to_f", [t])) |
|
146 |
handle SAME () => formula_from_atom (to_set Ts t) |
|
147 |
(* typ list -> term -> rel_expr *) |
|
148 |
and to_one Ts t = |
|
149 |
case t of |
|
150 |
Const (@{const_name Pair}, _) $ t1 $ t2 => |
|
151 |
Product (to_one Ts t1, to_one Ts t2) |
|
152 |
| Const (@{const_name Pair}, _) $ _ => to_one Ts (eta_expand Ts t 1) |
|
153 |
| Const (@{const_name Pair}, _) => to_one Ts (eta_expand Ts t 2) |
|
154 |
| Const (@{const_name fst}, _) $ t1 => |
|
155 |
let val fst_arity = arity_of_one scope (fastype_of1 (Ts, t)) in |
|
156 |
Project (to_one Ts t1, num_seq 0 fst_arity) |
|
157 |
end |
|
158 |
| Const (@{const_name fst}, _) => to_one Ts (eta_expand Ts t 1) |
|
159 |
| Const (@{const_name snd}, _) $ t1 => |
|
160 |
let |
|
161 |
val pair_arity = arity_of_one scope (fastype_of1 (Ts, t1)) |
|
162 |
val snd_arity = arity_of_one scope (fastype_of1 (Ts, t)) |
|
163 |
val fst_arity = pair_arity - snd_arity |
|
164 |
in Project (to_one Ts t1, num_seq fst_arity snd_arity) end |
|
165 |
| Const (@{const_name snd}, _) => to_one Ts (eta_expand Ts t 1) |
|
166 |
| Bound j => one_from_bound_var Ts j |
|
167 |
| _ => one_from_set Ts (fastype_of1 (Ts, t)) (to_set Ts t) |
|
168 |
(* term -> rel_expr *) |
|
169 |
and to_set Ts t = |
|
170 |
(case t of |
|
171 |
@{const Not} => to_set Ts (eta_expand Ts t 1) |
|
172 |
| Const (@{const_name All}, _) => to_set Ts (eta_expand Ts t 1) |
|
173 |
| Const (@{const_name Ex}, _) => to_set Ts (eta_expand Ts t 1) |
|
174 |
| Const (@{const_name "op ="}, _) $ _ => to_set Ts (eta_expand Ts t 1) |
|
175 |
| Const (@{const_name "op ="}, _) => to_set Ts (eta_expand Ts t 2) |
|
176 |
| Const (@{const_name ord_class.less_eq}, |
|
177 |
Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ _ => |
|
178 |
to_set Ts (eta_expand Ts t 1) |
|
179 |
| Const (@{const_name ord_class.less_eq}, _) => |
|
180 |
to_set Ts (eta_expand Ts t 2) |
|
181 |
| @{const "op &"} $ _ => to_set Ts (eta_expand Ts t 1) |
|
182 |
| @{const "op &"} => to_set Ts (eta_expand Ts t 2) |
|
183 |
| @{const "op |"} $ _ => to_set Ts (eta_expand Ts t 1) |
|
184 |
| @{const "op |"} => to_set Ts (eta_expand Ts t 2) |
|
185 |
| @{const "op -->"} $ _ => to_set Ts (eta_expand Ts t 1) |
|
186 |
| @{const "op -->"} => to_set Ts (eta_expand Ts t 2) |
|
187 |
| Const (@{const_name bot_class.bot}, |
|
188 |
T as Type ("fun", [_, @{typ bool}])) => |
|
189 |
empty_n_ary_rel (arity_of_set scope T) |
|
190 |
| Const (@{const_name insert}, _) $ t1 $ t2 => |
|
191 |
Union (to_one Ts t1, to_set Ts t2) |
|
192 |
| Const (@{const_name insert}, _) $ _ => to_set Ts (eta_expand Ts t 1) |
|
193 |
| Const (@{const_name insert}, _) => to_set Ts (eta_expand Ts t 2) |
|
194 |
| Const (@{const_name trancl}, _) $ t1 => |
|
195 |
if arity_of_set scope (fastype_of1 (Ts, t1)) = 2 then |
|
196 |
Closure (to_set Ts t1) |
|
197 |
else |
|
198 |
raise NOT_SUPPORTED "transitive closure for function or pair type" |
|
199 |
| Const (@{const_name trancl}, _) => to_set Ts (eta_expand Ts t 1) |
|
200 |
| Const (@{const_name lower_semilattice_class.inf}, |
|
201 |
Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ t1 $ t2 => |
|
202 |
Intersect (to_set Ts t1, to_set Ts t2) |
|
203 |
| Const (@{const_name lower_semilattice_class.inf}, _) $ _ => |
|
204 |
to_set Ts (eta_expand Ts t 1) |
|
205 |
| Const (@{const_name lower_semilattice_class.inf}, _) => |
|
206 |
to_set Ts (eta_expand Ts t 2) |
|
207 |
| Const (@{const_name upper_semilattice_class.sup}, |
|
208 |
Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ t1 $ t2 => |
|
209 |
Union (to_set Ts t1, to_set Ts t2) |
|
210 |
| Const (@{const_name upper_semilattice_class.sup}, _) $ _ => |
|
211 |
to_set Ts (eta_expand Ts t 1) |
|
212 |
| Const (@{const_name upper_semilattice_class.sup}, _) => |
|
213 |
to_set Ts (eta_expand Ts t 2) |
|
214 |
| Const (@{const_name minus_class.minus}, |
|
215 |
Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ t1 $ t2 => |
|
216 |
Difference (to_set Ts t1, to_set Ts t2) |
|
217 |
| Const (@{const_name minus_class.minus}, |
|
218 |
Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) $ _ => |
|
219 |
to_set Ts (eta_expand Ts t 1) |
|
220 |
| Const (@{const_name minus_class.minus}, |
|
221 |
Type ("fun", [Type ("fun", [_, @{typ bool}]), _])) => |
|
222 |
to_set Ts (eta_expand Ts t 2) |
|
223 |
| Const (@{const_name Pair}, _) $ _ $ _ => raise SAME () |
|
224 |
| Const (@{const_name Pair}, _) $ _ => raise SAME () |
|
225 |
| Const (@{const_name Pair}, _) => raise SAME () |
|
226 |
| Const (@{const_name fst}, _) $ _ => raise SAME () |
|
227 |
| Const (@{const_name fst}, _) => raise SAME () |
|
228 |
| Const (@{const_name snd}, _) $ _ => raise SAME () |
|
229 |
| Const (@{const_name snd}, _) => raise SAME () |
|
230 |
| Const (_, @{typ bool}) => atom_from_formula (to_f Ts t) |
|
231 |
| Free (x as (_, T)) => |
|
232 |
Rel (arity_of_set scope T, find_index (equal x) frees) |
|
233 |
| Term.Var _ => raise NOT_SUPPORTED "schematic variables" |
|
234 |
| Bound j => raise SAME () |
|
235 |
| Abs (_, T, t') => |
|
236 |
(case fastype_of1 (T :: Ts, t') of |
|
237 |
@{typ bool} => Comprehension (decls_for_one scope Ts T, |
|
238 |
to_f (T :: Ts) t') |
|
239 |
| T' => Comprehension (decls_for_one scope Ts T @ |
|
240 |
decls_for_set scope (T :: Ts) T', |
|
241 |
Subset (set_from_bound_var (T' :: T :: Ts) 0, |
|
242 |
to_set (T :: Ts) t'))) |
|
243 |
| t1 $ t2 => |
|
244 |
(case fastype_of1 (Ts, t) of |
|
245 |
@{typ bool} => atom_from_formula (to_f Ts t) |
|
246 |
| T => |
|
247 |
let val T2 = fastype_of1 (Ts, t2) in |
|
248 |
case arity_of_one scope T2 of |
|
249 |
1 => Join (to_one Ts t2, to_set Ts t1) |
|
250 |
| n => |
|
251 |
let |
|
252 |
val arity2 = arity_of_one scope T2 |
|
253 |
val res_arity = arity_of_set scope T |
|
254 |
in |
|
255 |
Project (Intersect |
|
256 |
(Product (to_one Ts t2, |
|
257 |
atom_schema_of_set scope T |
|
258 |
|> map (AtomSeq o rpair 0) |> foldl1 Product), |
|
259 |
to_set Ts t1), |
|
260 |
num_seq arity2 res_arity) |
|
261 |
end |
|
262 |
end) |
|
263 |
| _ => raise NOT_SUPPORTED ("term " ^ |
|
264 |
quote (Syntax.string_of_term ctxt t))) |
|
265 |
handle SAME () => set_from_one (fastype_of1 (Ts, t)) (to_one Ts t) |
|
266 |
in to_f [] end |
|
267 |
||
268 |
(* (typ -> int) -> int -> styp -> bound *) |
|
269 |
fun bound_for_free scope i (s, T) = |
|
270 |
let val js = atom_schema_of_set scope T in |
|
271 |
([((length js, i), s)], |
|
272 |
[TupleSet [], atom_schema_of_set scope T |> map (rpair 0) |
|
273 |
|> tuple_set_from_atom_schema]) |
|
274 |
end |
|
275 |
||
276 |
(* (typ -> int) -> typ list -> typ -> rel_expr -> formula *) |
|
277 |
fun declarative_axiom_for_rel_expr scope Ts (Type ("fun", [T1, T2])) r = |
|
278 |
if body_type T2 = bool_T then |
|
279 |
True |
|
280 |
else |
|
281 |
All (decls_for_one scope Ts T1, |
|
282 |
declarative_axiom_for_rel_expr scope (T1 :: Ts) T2 |
|
283 |
(List.foldl Join r (one_vars_for_bound_var scope (T1 :: Ts) 0))) |
|
284 |
| declarative_axiom_for_rel_expr _ _ _ r = One r |
|
285 |
||
286 |
(* (typ -> int) -> int -> styp -> formula *) |
|
287 |
fun declarative_axiom_for_free scope i (_, T) = |
|
288 |
declarative_axiom_for_rel_expr scope [] T (Rel (arity_of_set scope T, i)) |
|
289 |
||
290 |
(* Proof.context -> (typ -> int) -> term -> string *) |
|
291 |
fun pick_nits_in_term ctxt raw_scope t = |
|
292 |
let |
|
293 |
val thy = ProofContext.theory_of ctxt |
|
294 |
(* typ -> int *) |
|
295 |
fun scope (Type ("fun", [T1, T2])) = reasonable_power (scope T2) (scope T1) |
|
296 |
| scope (Type ("*", [T1, T2])) = scope T1 * scope T2 |
|
297 |
| scope @{typ bool} = 2 |
|
298 |
| scope T = Int.max (1, raw_scope T) |
|
299 |
val neg_t = @{const Not} $ ObjectLogic.atomize_term thy t |
|
300 |
val _ = fold_types (K o check_type ctxt) neg_t () |
|
301 |
val frees = Term.add_frees neg_t [] |
|
302 |
val bounds = map2 (bound_for_free scope) (index_seq 0 (length frees)) frees |
|
303 |
val declarative_axioms = |
|
304 |
map2 (declarative_axiom_for_free scope) (index_seq 0 (length frees)) |
|
305 |
frees |
|
306 |
val formula = kodkod_formula_for_term ctxt scope frees neg_t |
|
307 |
|> fold_rev (curry And) declarative_axioms |
|
308 |
val univ_card = univ_card 0 0 0 bounds formula |
|
309 |
val problem = |
|
310 |
{comment = "", settings = [], univ_card = univ_card, tuple_assigns = [], |
|
311 |
bounds = bounds, int_bounds = [], expr_assigns = [], formula = formula} |
|
312 |
in |
|
313 |
case solve_any_problem true NONE 0 1 [problem] of |
|
314 |
Normal ([], _) => "none" |
|
315 |
| Normal _ => "genuine" |
|
316 |
| TimedOut _ => "unknown" |
|
317 |
| Interrupted _ => "unknown" |
|
318 |
| Error (s, _) => error ("Kodkod error: " ^ s) |
|
319 |
end |
|
320 |
handle NOT_SUPPORTED details => |
|
321 |
(warning ("Unsupported case: " ^ details ^ "."); "unknown") |
|
322 |
end; |