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