author | blanchet |
Mon, 22 Feb 2010 11:57:33 +0100 | |
changeset 35280 | 54ab4921f826 |
parent 35220 | 2bcdae5f4fdb |
child 35384 | 88dbcfe75c45 |
permissions | -rw-r--r-- |
33982 | 1 |
(* Title: HOL/Tools/Nitpick/nitpick_mono.ML |
33192 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
3 |
Copyright 2009, 2010 |
33192 | 4 |
|
5 |
Monotonicity predicate for higher-order logic. |
|
6 |
*) |
|
7 |
||
8 |
signature NITPICK_MONO = |
|
9 |
sig |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
10 |
datatype sign = Plus | Minus |
35070
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents:
34998
diff
changeset
|
11 |
type hol_context = Nitpick_HOL.hol_context |
33192 | 12 |
|
13 |
val formulas_monotonic : |
|
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
14 |
hol_context -> bool -> typ -> sign -> term list -> term list -> term -> bool |
33192 | 15 |
end; |
16 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
17 |
structure Nitpick_Mono : NITPICK_MONO = |
33192 | 18 |
struct |
19 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
20 |
open Nitpick_Util |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
21 |
open Nitpick_HOL |
33192 | 22 |
|
23 |
type var = int |
|
24 |
||
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
25 |
datatype sign = Plus | Minus |
33192 | 26 |
datatype sign_atom = S of sign | V of var |
27 |
||
28 |
type literal = var * sign |
|
29 |
||
30 |
datatype ctype = |
|
31 |
CAlpha | |
|
32 |
CFun of ctype * sign_atom * ctype | |
|
33 |
CPair of ctype * ctype | |
|
34 |
CType of string * ctype list | |
|
35 |
CRec of string * typ list |
|
36 |
||
37 |
type cdata = |
|
35070
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents:
34998
diff
changeset
|
38 |
{hol_ctxt: hol_context, |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
39 |
binarize: bool, |
33192 | 40 |
alpha_T: typ, |
41 |
max_fresh: int Unsynchronized.ref, |
|
42 |
datatype_cache: ((string * typ list) * ctype) list Unsynchronized.ref, |
|
43 |
constr_cache: (styp * ctype) list Unsynchronized.ref} |
|
44 |
||
45 |
exception CTYPE of string * ctype list |
|
46 |
||
47 |
(* string -> unit *) |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
48 |
fun print_g (_ : string) = () |
33192 | 49 |
|
50 |
(* var -> string *) |
|
51 |
val string_for_var = signed_string_of_int |
|
52 |
(* string -> var list -> string *) |
|
53 |
fun string_for_vars sep [] = "0\<^bsub>" ^ sep ^ "\<^esub>" |
|
54 |
| string_for_vars sep xs = space_implode sep (map string_for_var xs) |
|
55 |
fun subscript_string_for_vars sep xs = |
|
56 |
if null xs then "" else "\<^bsub>" ^ string_for_vars sep xs ^ "\<^esub>" |
|
57 |
||
58 |
(* sign -> string *) |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
59 |
fun string_for_sign Plus = "+" |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
60 |
| string_for_sign Minus = "-" |
33192 | 61 |
|
62 |
(* sign -> sign -> sign *) |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
63 |
fun xor sn1 sn2 = if sn1 = sn2 then Plus else Minus |
33192 | 64 |
(* sign -> sign *) |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
65 |
val negate = xor Minus |
33192 | 66 |
|
67 |
(* sign_atom -> string *) |
|
68 |
fun string_for_sign_atom (S sn) = string_for_sign sn |
|
69 |
| string_for_sign_atom (V j) = string_for_var j |
|
70 |
||
71 |
(* literal -> string *) |
|
72 |
fun string_for_literal (x, sn) = string_for_var x ^ " = " ^ string_for_sign sn |
|
73 |
||
74 |
val bool_C = CType (@{type_name bool}, []) |
|
75 |
||
76 |
(* ctype -> bool *) |
|
77 |
fun is_CRec (CRec _) = true |
|
78 |
| is_CRec _ = false |
|
79 |
||
80 |
val no_prec = 100 |
|
81 |
val prec_CFun = 1 |
|
82 |
val prec_CPair = 2 |
|
83 |
||
84 |
(* tuple_set -> int *) |
|
85 |
fun precedence_of_ctype (CFun _) = prec_CFun |
|
86 |
| precedence_of_ctype (CPair _) = prec_CPair |
|
87 |
| precedence_of_ctype _ = no_prec |
|
88 |
||
89 |
(* ctype -> string *) |
|
90 |
val string_for_ctype = |
|
91 |
let |
|
92 |
(* int -> ctype -> string *) |
|
93 |
fun aux outer_prec C = |
|
94 |
let |
|
95 |
val prec = precedence_of_ctype C |
|
96 |
val need_parens = (prec < outer_prec) |
|
97 |
in |
|
98 |
(if need_parens then "(" else "") ^ |
|
99 |
(case C of |
|
100 |
CAlpha => "\<alpha>" |
|
101 |
| CFun (C1, a, C2) => |
|
102 |
aux (prec + 1) C1 ^ " \<Rightarrow>\<^bsup>" ^ |
|
103 |
string_for_sign_atom a ^ "\<^esup> " ^ aux prec C2 |
|
104 |
| CPair (C1, C2) => aux (prec + 1) C1 ^ " \<times> " ^ aux prec C2 |
|
105 |
| CType (s, []) => |
|
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
106 |
if s = @{type_name prop} orelse s = @{type_name bool} then "o" else s |
33192 | 107 |
| CType (s, Cs) => "(" ^ commas (map (aux 0) Cs) ^ ") " ^ s |
108 |
| CRec (s, _) => "[" ^ s ^ "]") ^ |
|
109 |
(if need_parens then ")" else "") |
|
110 |
end |
|
111 |
in aux 0 end |
|
112 |
||
113 |
(* ctype -> ctype list *) |
|
114 |
fun flatten_ctype (CPair (C1, C2)) = maps flatten_ctype [C1, C2] |
|
115 |
| flatten_ctype (CType (_, Cs)) = maps flatten_ctype Cs |
|
116 |
| flatten_ctype C = [C] |
|
117 |
||
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
118 |
(* hol_context -> bool -> typ -> cdata *) |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
119 |
fun initial_cdata hol_ctxt binarize alpha_T = |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
120 |
({hol_ctxt = hol_ctxt, binarize = binarize, alpha_T = alpha_T, |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
121 |
max_fresh = Unsynchronized.ref 0, datatype_cache = Unsynchronized.ref [], |
33192 | 122 |
constr_cache = Unsynchronized.ref []} : cdata) |
123 |
||
124 |
(* typ -> typ -> bool *) |
|
125 |
fun could_exist_alpha_subtype alpha_T (T as Type (_, Ts)) = |
|
34936
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34123
diff
changeset
|
126 |
T = alpha_T orelse (not (is_fp_iterator_type T) andalso |
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34123
diff
changeset
|
127 |
exists (could_exist_alpha_subtype alpha_T) Ts) |
33192 | 128 |
| could_exist_alpha_subtype alpha_T T = (T = alpha_T) |
129 |
(* theory -> typ -> typ -> bool *) |
|
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
130 |
fun could_exist_alpha_sub_ctype _ (alpha_T as TFree _) T = |
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
131 |
could_exist_alpha_subtype alpha_T T |
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
132 |
| could_exist_alpha_sub_ctype thy alpha_T T = |
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
133 |
(T = alpha_T orelse is_datatype thy [(NONE, true)] T) |
33192 | 134 |
|
135 |
(* ctype -> bool *) |
|
136 |
fun exists_alpha_sub_ctype CAlpha = true |
|
137 |
| exists_alpha_sub_ctype (CFun (C1, _, C2)) = |
|
138 |
exists exists_alpha_sub_ctype [C1, C2] |
|
139 |
| exists_alpha_sub_ctype (CPair (C1, C2)) = |
|
140 |
exists exists_alpha_sub_ctype [C1, C2] |
|
141 |
| exists_alpha_sub_ctype (CType (_, Cs)) = exists exists_alpha_sub_ctype Cs |
|
142 |
| exists_alpha_sub_ctype (CRec _) = true |
|
143 |
||
144 |
(* ctype -> bool *) |
|
145 |
fun exists_alpha_sub_ctype_fresh CAlpha = true |
|
146 |
| exists_alpha_sub_ctype_fresh (CFun (_, V _, _)) = true |
|
147 |
| exists_alpha_sub_ctype_fresh (CFun (_, _, C2)) = |
|
148 |
exists_alpha_sub_ctype_fresh C2 |
|
149 |
| exists_alpha_sub_ctype_fresh (CPair (C1, C2)) = |
|
150 |
exists exists_alpha_sub_ctype_fresh [C1, C2] |
|
151 |
| exists_alpha_sub_ctype_fresh (CType (_, Cs)) = |
|
152 |
exists exists_alpha_sub_ctype_fresh Cs |
|
153 |
| exists_alpha_sub_ctype_fresh (CRec _) = true |
|
154 |
||
155 |
(* string * typ list -> ctype list -> ctype *) |
|
156 |
fun constr_ctype_for_binders z Cs = |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
157 |
fold_rev (fn C => curry3 CFun C (S Minus)) Cs (CRec z) |
33192 | 158 |
|
159 |
(* ((string * typ list) * ctype) list -> ctype list -> ctype -> ctype *) |
|
160 |
fun repair_ctype _ _ CAlpha = CAlpha |
|
161 |
| repair_ctype cache seen (CFun (C1, a, C2)) = |
|
162 |
CFun (repair_ctype cache seen C1, a, repair_ctype cache seen C2) |
|
163 |
| repair_ctype cache seen (CPair Cp) = |
|
164 |
CPair (pairself (repair_ctype cache seen) Cp) |
|
165 |
| repair_ctype cache seen (CType (s, Cs)) = |
|
166 |
CType (s, maps (flatten_ctype o repair_ctype cache seen) Cs) |
|
167 |
| repair_ctype cache seen (CRec (z as (s, _))) = |
|
168 |
case AList.lookup (op =) cache z |> the of |
|
169 |
CRec _ => CType (s, []) |
|
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
170 |
| C => if member (op =) seen C then CType (s, []) |
33192 | 171 |
else repair_ctype cache (C :: seen) C |
172 |
||
173 |
(* ((string * typ list) * ctype) list Unsynchronized.ref -> unit *) |
|
174 |
fun repair_datatype_cache cache = |
|
175 |
let |
|
176 |
(* (string * typ list) * ctype -> unit *) |
|
177 |
fun repair_one (z, C) = |
|
178 |
Unsynchronized.change cache |
|
179 |
(AList.update (op =) (z, repair_ctype (!cache) [] C)) |
|
180 |
in List.app repair_one (rev (!cache)) end |
|
181 |
||
182 |
(* (typ * ctype) list -> (styp * ctype) list Unsynchronized.ref -> unit *) |
|
183 |
fun repair_constr_cache dtype_cache constr_cache = |
|
184 |
let |
|
185 |
(* styp * ctype -> unit *) |
|
186 |
fun repair_one (x, C) = |
|
187 |
Unsynchronized.change constr_cache |
|
188 |
(AList.update (op =) (x, repair_ctype dtype_cache [] C)) |
|
189 |
in List.app repair_one (!constr_cache) end |
|
190 |
||
191 |
(* cdata -> typ -> ctype *) |
|
35219
15a5f213ef5b
fix bug in Nitpick's monotonicity code w.r.t. binary integers
blanchet
parents:
35190
diff
changeset
|
192 |
fun fresh_ctype_for_type ({hol_ctxt as {thy, ...}, binarize, alpha_T, max_fresh, |
33192 | 193 |
datatype_cache, constr_cache, ...} : cdata) = |
194 |
let |
|
195 |
(* typ -> typ -> ctype *) |
|
196 |
fun do_fun T1 T2 = |
|
197 |
let |
|
198 |
val C1 = do_type T1 |
|
199 |
val C2 = do_type T2 |
|
34936
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34123
diff
changeset
|
200 |
val a = if is_boolean_type (body_type T2) andalso |
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34123
diff
changeset
|
201 |
exists_alpha_sub_ctype_fresh C1 then |
33192 | 202 |
V (Unsynchronized.inc max_fresh) |
203 |
else |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
204 |
S Minus |
33192 | 205 |
in CFun (C1, a, C2) end |
206 |
(* typ -> ctype *) |
|
207 |
and do_type T = |
|
208 |
if T = alpha_T then |
|
209 |
CAlpha |
|
210 |
else case T of |
|
211 |
Type ("fun", [T1, T2]) => do_fun T1 T2 |
|
212 |
| Type (@{type_name fun_box}, [T1, T2]) => do_fun T1 T2 |
|
213 |
| Type ("*", [T1, T2]) => CPair (pairself do_type (T1, T2)) |
|
214 |
| Type (z as (s, _)) => |
|
215 |
if could_exist_alpha_sub_ctype thy alpha_T T then |
|
216 |
case AList.lookup (op =) (!datatype_cache) z of |
|
217 |
SOME C => C |
|
218 |
| NONE => |
|
219 |
let |
|
220 |
val _ = Unsynchronized.change datatype_cache (cons (z, CRec z)) |
|
35219
15a5f213ef5b
fix bug in Nitpick's monotonicity code w.r.t. binary integers
blanchet
parents:
35190
diff
changeset
|
221 |
val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T |
33192 | 222 |
val (all_Cs, constr_Cs) = |
223 |
fold_rev (fn (_, T') => fn (all_Cs, constr_Cs) => |
|
224 |
let |
|
225 |
val binder_Cs = map do_type (binder_types T') |
|
226 |
val new_Cs = filter exists_alpha_sub_ctype_fresh |
|
227 |
binder_Cs |
|
228 |
val constr_C = constr_ctype_for_binders z |
|
229 |
binder_Cs |
|
230 |
in |
|
231 |
(union (op =) new_Cs all_Cs, |
|
232 |
constr_C :: constr_Cs) |
|
233 |
end) |
|
234 |
xs ([], []) |
|
235 |
val C = CType (s, all_Cs) |
|
236 |
val _ = Unsynchronized.change datatype_cache |
|
237 |
(AList.update (op =) (z, C)) |
|
238 |
val _ = Unsynchronized.change constr_cache |
|
239 |
(append (xs ~~ constr_Cs)) |
|
240 |
in |
|
241 |
if forall (not o is_CRec o snd) (!datatype_cache) then |
|
242 |
(repair_datatype_cache datatype_cache; |
|
243 |
repair_constr_cache (!datatype_cache) constr_cache; |
|
244 |
AList.lookup (op =) (!datatype_cache) z |> the) |
|
245 |
else |
|
246 |
C |
|
247 |
end |
|
248 |
else |
|
249 |
CType (s, []) |
|
250 |
| _ => CType (Refute.string_of_typ T, []) |
|
251 |
in do_type end |
|
252 |
||
253 |
(* ctype -> ctype list *) |
|
254 |
fun prodC_factors (CPair (C1, C2)) = maps prodC_factors [C1, C2] |
|
255 |
| prodC_factors C = [C] |
|
256 |
(* ctype -> ctype list * ctype *) |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
257 |
fun curried_strip_ctype (CFun (C1, S Minus, C2)) = |
33192 | 258 |
curried_strip_ctype C2 |>> append (prodC_factors C1) |
259 |
| curried_strip_ctype C = ([], C) |
|
260 |
(* string -> ctype -> ctype *) |
|
261 |
fun sel_ctype_from_constr_ctype s C = |
|
262 |
let val (arg_Cs, dataC) = curried_strip_ctype C in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
263 |
CFun (dataC, S Minus, |
33192 | 264 |
case sel_no_from_name s of ~1 => bool_C | n => nth arg_Cs n) |
265 |
end |
|
266 |
||
267 |
(* cdata -> styp -> ctype *) |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
268 |
fun ctype_for_constr (cdata as {hol_ctxt = {thy, ...}, alpha_T, constr_cache, |
33192 | 269 |
...}) (x as (_, T)) = |
270 |
if could_exist_alpha_sub_ctype thy alpha_T T then |
|
271 |
case AList.lookup (op =) (!constr_cache) x of |
|
272 |
SOME C => C |
|
35219
15a5f213ef5b
fix bug in Nitpick's monotonicity code w.r.t. binary integers
blanchet
parents:
35190
diff
changeset
|
273 |
| NONE => (fresh_ctype_for_type cdata (body_type T); |
15a5f213ef5b
fix bug in Nitpick's monotonicity code w.r.t. binary integers
blanchet
parents:
35190
diff
changeset
|
274 |
AList.lookup (op =) (!constr_cache) x |> the) |
33192 | 275 |
else |
276 |
fresh_ctype_for_type cdata T |
|
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
277 |
fun ctype_for_sel (cdata as {hol_ctxt, binarize, ...}) (x as (s, _)) = |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
278 |
x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
279 |
|> ctype_for_constr cdata |> sel_ctype_from_constr_ctype s |
33192 | 280 |
|
281 |
(* literal list -> ctype -> ctype *) |
|
282 |
fun instantiate_ctype lits = |
|
283 |
let |
|
284 |
(* ctype -> ctype *) |
|
285 |
fun aux CAlpha = CAlpha |
|
286 |
| aux (CFun (C1, V x, C2)) = |
|
287 |
let |
|
288 |
val a = case AList.lookup (op =) lits x of |
|
289 |
SOME sn => S sn |
|
290 |
| NONE => V x |
|
291 |
in CFun (aux C1, a, aux C2) end |
|
292 |
| aux (CFun (C1, a, C2)) = CFun (aux C1, a, aux C2) |
|
293 |
| aux (CPair Cp) = CPair (pairself aux Cp) |
|
294 |
| aux (CType (s, Cs)) = CType (s, map aux Cs) |
|
295 |
| aux (CRec z) = CRec z |
|
296 |
in aux end |
|
297 |
||
298 |
datatype comp_op = Eq | Leq |
|
299 |
||
300 |
type comp = sign_atom * sign_atom * comp_op * var list |
|
301 |
type sign_expr = literal list |
|
302 |
||
303 |
datatype constraint_set = |
|
304 |
UnsolvableCSet | |
|
305 |
CSet of literal list * comp list * sign_expr list |
|
306 |
||
307 |
(* comp_op -> string *) |
|
308 |
fun string_for_comp_op Eq = "=" |
|
309 |
| string_for_comp_op Leq = "\<le>" |
|
310 |
||
311 |
(* sign_expr -> string *) |
|
312 |
fun string_for_sign_expr [] = "\<bot>" |
|
313 |
| string_for_sign_expr lits = |
|
314 |
space_implode " \<or> " (map string_for_literal lits) |
|
315 |
||
316 |
(* constraint_set *) |
|
317 |
val slack = CSet ([], [], []) |
|
318 |
||
319 |
(* literal -> literal list option -> literal list option *) |
|
320 |
fun do_literal _ NONE = NONE |
|
321 |
| do_literal (x, sn) (SOME lits) = |
|
322 |
case AList.lookup (op =) lits x of |
|
323 |
SOME sn' => if sn = sn' then SOME lits else NONE |
|
324 |
| NONE => SOME ((x, sn) :: lits) |
|
325 |
||
326 |
(* comp_op -> var list -> sign_atom -> sign_atom -> literal list * comp list |
|
327 |
-> (literal list * comp list) option *) |
|
328 |
fun do_sign_atom_comp Eq [] a1 a2 (accum as (lits, comps)) = |
|
329 |
(case (a1, a2) of |
|
330 |
(S sn1, S sn2) => if sn1 = sn2 then SOME accum else NONE |
|
331 |
| (V x1, S sn2) => |
|
332 |
Option.map (rpair comps) (do_literal (x1, sn2) (SOME lits)) |
|
333 |
| (V _, V _) => SOME (lits, insert (op =) (a1, a2, Eq, []) comps) |
|
334 |
| _ => do_sign_atom_comp Eq [] a2 a1 accum) |
|
335 |
| do_sign_atom_comp Leq [] a1 a2 (accum as (lits, comps)) = |
|
336 |
(case (a1, a2) of |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
337 |
(_, S Minus) => SOME accum |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
338 |
| (S Plus, _) => SOME accum |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
339 |
| (S Minus, S Plus) => NONE |
33192 | 340 |
| (V _, V _) => SOME (lits, insert (op =) (a1, a2, Leq, []) comps) |
341 |
| _ => do_sign_atom_comp Eq [] a1 a2 accum) |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
342 |
| do_sign_atom_comp cmp xs a1 a2 (lits, comps) = |
33192 | 343 |
SOME (lits, insert (op =) (a1, a2, cmp, xs) comps) |
344 |
||
345 |
(* comp -> var list -> ctype -> ctype -> (literal list * comp list) option |
|
346 |
-> (literal list * comp list) option *) |
|
347 |
fun do_ctype_comp _ _ _ _ NONE = NONE |
|
348 |
| do_ctype_comp _ _ CAlpha CAlpha accum = accum |
|
349 |
| do_ctype_comp Eq xs (CFun (C11, a1, C12)) (CFun (C21, a2, C22)) |
|
350 |
(SOME accum) = |
|
351 |
accum |> do_sign_atom_comp Eq xs a1 a2 |> do_ctype_comp Eq xs C11 C21 |
|
352 |
|> do_ctype_comp Eq xs C12 C22 |
|
353 |
| do_ctype_comp Leq xs (CFun (C11, a1, C12)) (CFun (C21, a2, C22)) |
|
354 |
(SOME accum) = |
|
355 |
(if exists_alpha_sub_ctype C11 then |
|
356 |
accum |> do_sign_atom_comp Leq xs a1 a2 |
|
357 |
|> do_ctype_comp Leq xs C21 C11 |
|
358 |
|> (case a2 of |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
359 |
S Minus => I |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
360 |
| S Plus => do_ctype_comp Leq xs C11 C21 |
33192 | 361 |
| V x => do_ctype_comp Leq (x :: xs) C11 C21) |
362 |
else |
|
363 |
SOME accum) |
|
364 |
|> do_ctype_comp Leq xs C12 C22 |
|
365 |
| do_ctype_comp cmp xs (C1 as CPair (C11, C12)) (C2 as CPair (C21, C22)) |
|
366 |
accum = |
|
367 |
(accum |> fold (uncurry (do_ctype_comp cmp xs)) [(C11, C21), (C12, C22)] |
|
368 |
handle Library.UnequalLengths => |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
369 |
raise CTYPE ("Nitpick_Mono.do_ctype_comp", [C1, C2])) |
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
370 |
| do_ctype_comp _ _ (CType _) (CType _) accum = |
33192 | 371 |
accum (* no need to compare them thanks to the cache *) |
372 |
| do_ctype_comp _ _ C1 C2 _ = |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
373 |
raise CTYPE ("Nitpick_Mono.do_ctype_comp", [C1, C2]) |
33192 | 374 |
|
375 |
(* comp_op -> ctype -> ctype -> constraint_set -> constraint_set *) |
|
376 |
fun add_ctype_comp _ _ _ UnsolvableCSet = UnsolvableCSet |
|
377 |
| add_ctype_comp cmp C1 C2 (CSet (lits, comps, sexps)) = |
|
378 |
(print_g ("*** Add " ^ string_for_ctype C1 ^ " " ^ string_for_comp_op cmp ^ |
|
379 |
" " ^ string_for_ctype C2); |
|
380 |
case do_ctype_comp cmp [] C1 C2 (SOME (lits, comps)) of |
|
381 |
NONE => (print_g "**** Unsolvable"; UnsolvableCSet) |
|
382 |
| SOME (lits, comps) => CSet (lits, comps, sexps)) |
|
383 |
||
384 |
(* ctype -> ctype -> constraint_set -> constraint_set *) |
|
385 |
val add_ctypes_equal = add_ctype_comp Eq |
|
386 |
val add_is_sub_ctype = add_ctype_comp Leq |
|
387 |
||
388 |
(* sign -> sign_expr -> ctype -> (literal list * sign_expr list) option |
|
389 |
-> (literal list * sign_expr list) option *) |
|
390 |
fun do_notin_ctype_fv _ _ _ NONE = NONE |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
391 |
| do_notin_ctype_fv Minus _ CAlpha accum = accum |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
392 |
| do_notin_ctype_fv Plus [] CAlpha _ = NONE |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
393 |
| do_notin_ctype_fv Plus [(x, sn)] CAlpha (SOME (lits, sexps)) = |
33192 | 394 |
SOME lits |> do_literal (x, sn) |> Option.map (rpair sexps) |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
395 |
| do_notin_ctype_fv Plus sexp CAlpha (SOME (lits, sexps)) = |
33192 | 396 |
SOME (lits, insert (op =) sexp sexps) |
397 |
| do_notin_ctype_fv sn sexp (CFun (C1, S sn', C2)) accum = |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
398 |
accum |> (if sn' = Plus andalso sn = Plus then |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
399 |
do_notin_ctype_fv Plus sexp C1 |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
400 |
else |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
401 |
I) |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
402 |
|> (if sn' = Minus orelse sn = Plus then |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
403 |
do_notin_ctype_fv Minus sexp C1 |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
404 |
else |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
405 |
I) |
33192 | 406 |
|> do_notin_ctype_fv sn sexp C2 |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
407 |
| do_notin_ctype_fv Plus sexp (CFun (C1, V x, C2)) accum = |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
408 |
accum |> (case do_literal (x, Minus) (SOME sexp) of |
33192 | 409 |
NONE => I |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
410 |
| SOME sexp' => do_notin_ctype_fv Plus sexp' C1) |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
411 |
|> do_notin_ctype_fv Minus sexp C1 |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
412 |
|> do_notin_ctype_fv Plus sexp C2 |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
413 |
| do_notin_ctype_fv Minus sexp (CFun (C1, V x, C2)) accum = |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
414 |
accum |> (case do_literal (x, Plus) (SOME sexp) of |
33192 | 415 |
NONE => I |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
416 |
| SOME sexp' => do_notin_ctype_fv Plus sexp' C1) |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
417 |
|> do_notin_ctype_fv Minus sexp C2 |
33192 | 418 |
| do_notin_ctype_fv sn sexp (CPair (C1, C2)) accum = |
419 |
accum |> fold (do_notin_ctype_fv sn sexp) [C1, C2] |
|
420 |
| do_notin_ctype_fv sn sexp (CType (_, Cs)) accum = |
|
421 |
accum |> fold (do_notin_ctype_fv sn sexp) Cs |
|
422 |
| do_notin_ctype_fv _ _ C _ = |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
423 |
raise CTYPE ("Nitpick_Mono.do_notin_ctype_fv", [C]) |
33192 | 424 |
|
425 |
(* sign -> ctype -> constraint_set -> constraint_set *) |
|
426 |
fun add_notin_ctype_fv _ _ UnsolvableCSet = UnsolvableCSet |
|
427 |
| add_notin_ctype_fv sn C (CSet (lits, comps, sexps)) = |
|
428 |
(print_g ("*** Add " ^ string_for_ctype C ^ " is right-" ^ |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
429 |
(case sn of Minus => "unique" | Plus => "total") ^ "."); |
33192 | 430 |
case do_notin_ctype_fv sn [] C (SOME (lits, sexps)) of |
431 |
NONE => (print_g "**** Unsolvable"; UnsolvableCSet) |
|
432 |
| SOME (lits, sexps) => CSet (lits, comps, sexps)) |
|
433 |
||
434 |
(* ctype -> constraint_set -> constraint_set *) |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
435 |
val add_ctype_is_right_unique = add_notin_ctype_fv Minus |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
436 |
val add_ctype_is_right_total = add_notin_ctype_fv Plus |
33192 | 437 |
|
438 |
(* sign -> bool *) |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
439 |
fun bool_from_sign Plus = false |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
440 |
| bool_from_sign Minus = true |
33192 | 441 |
(* bool -> sign *) |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
442 |
fun sign_from_bool false = Plus |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
443 |
| sign_from_bool true = Minus |
33192 | 444 |
|
445 |
(* literal -> PropLogic.prop_formula *) |
|
446 |
fun prop_for_literal (x, sn) = |
|
447 |
(not (bool_from_sign sn) ? PropLogic.Not) (PropLogic.BoolVar x) |
|
448 |
(* sign_atom -> PropLogic.prop_formula *) |
|
449 |
fun prop_for_sign_atom_eq (S sn', sn) = |
|
450 |
if sn = sn' then PropLogic.True else PropLogic.False |
|
451 |
| prop_for_sign_atom_eq (V x, sn) = prop_for_literal (x, sn) |
|
452 |
(* sign_expr -> PropLogic.prop_formula *) |
|
453 |
fun prop_for_sign_expr xs = PropLogic.exists (map prop_for_literal xs) |
|
454 |
(* var list -> sign -> PropLogic.prop_formula *) |
|
455 |
fun prop_for_exists_eq xs sn = |
|
456 |
PropLogic.exists (map (fn x => prop_for_literal (x, sn)) xs) |
|
457 |
(* comp -> PropLogic.prop_formula *) |
|
458 |
fun prop_for_comp (a1, a2, Eq, []) = |
|
459 |
PropLogic.SAnd (prop_for_comp (a1, a2, Leq, []), |
|
460 |
prop_for_comp (a2, a1, Leq, [])) |
|
461 |
| prop_for_comp (a1, a2, Leq, []) = |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
462 |
PropLogic.SOr (prop_for_sign_atom_eq (a1, Plus), |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
463 |
prop_for_sign_atom_eq (a2, Minus)) |
33192 | 464 |
| prop_for_comp (a1, a2, cmp, xs) = |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
465 |
PropLogic.SOr (prop_for_exists_eq xs Minus, prop_for_comp (a1, a2, cmp, [])) |
33192 | 466 |
|
467 |
(* var -> (int -> bool option) -> literal list -> literal list *) |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
468 |
fun literals_from_assignments max_var assigns lits = |
33192 | 469 |
fold (fn x => fn accum => |
470 |
if AList.defined (op =) lits x then |
|
471 |
accum |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
472 |
else case assigns x of |
33192 | 473 |
SOME b => (x, sign_from_bool b) :: accum |
474 |
| NONE => accum) (max_var downto 1) lits |
|
475 |
||
476 |
(* comp -> string *) |
|
477 |
fun string_for_comp (a1, a2, cmp, xs) = |
|
478 |
string_for_sign_atom a1 ^ " " ^ string_for_comp_op cmp ^ |
|
479 |
subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_sign_atom a2 |
|
480 |
||
481 |
(* literal list -> comp list -> sign_expr list -> unit *) |
|
482 |
fun print_problem lits comps sexps = |
|
483 |
print_g ("*** Problem:\n" ^ cat_lines (map string_for_literal lits @ |
|
484 |
map string_for_comp comps @ |
|
485 |
map string_for_sign_expr sexps)) |
|
486 |
||
487 |
(* literal list -> unit *) |
|
488 |
fun print_solution lits = |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
489 |
let val (pos, neg) = List.partition (curry (op =) Plus o snd) lits in |
33192 | 490 |
print_g ("*** Solution:\n" ^ |
491 |
"+: " ^ commas (map (string_for_var o fst) pos) ^ "\n" ^ |
|
492 |
"-: " ^ commas (map (string_for_var o fst) neg)) |
|
493 |
end |
|
494 |
||
495 |
(* var -> constraint_set -> literal list list option *) |
|
496 |
fun solve _ UnsolvableCSet = (print_g "*** Problem: Unsolvable"; NONE) |
|
497 |
| solve max_var (CSet (lits, comps, sexps)) = |
|
498 |
let |
|
499 |
val _ = print_problem lits comps sexps |
|
500 |
val prop = PropLogic.all (map prop_for_literal lits @ |
|
501 |
map prop_for_comp comps @ |
|
502 |
map prop_for_sign_expr sexps) |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
503 |
(* use the first ML solver (to avoid startup overhead) *) |
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
504 |
val solvers = !SatSolver.solvers |
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
505 |
|> filter (member (op =) ["dptsat", "dpll"] o fst) |
33192 | 506 |
in |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
507 |
case snd (hd solvers) prop of |
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
508 |
SatSolver.SATISFIABLE assigns => |
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
509 |
SOME (literals_from_assignments max_var assigns lits |
33192 | 510 |
|> tap print_solution) |
511 |
| _ => NONE |
|
512 |
end |
|
513 |
||
514 |
type ctype_schema = ctype * constraint_set |
|
515 |
type ctype_context = |
|
516 |
{bounds: ctype list, |
|
517 |
frees: (styp * ctype) list, |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
518 |
consts: (styp * ctype) list} |
33192 | 519 |
|
520 |
type accumulator = ctype_context * constraint_set |
|
521 |
||
522 |
val initial_gamma = {bounds = [], frees = [], consts = []} |
|
523 |
val unsolvable_accum = (initial_gamma, UnsolvableCSet) |
|
524 |
||
525 |
(* ctype -> ctype_context -> ctype_context *) |
|
526 |
fun push_bound C {bounds, frees, consts} = |
|
527 |
{bounds = C :: bounds, frees = frees, consts = consts} |
|
528 |
(* ctype_context -> ctype_context *) |
|
529 |
fun pop_bound {bounds, frees, consts} = |
|
530 |
{bounds = tl bounds, frees = frees, consts = consts} |
|
531 |
handle List.Empty => initial_gamma |
|
532 |
||
533 |
(* cdata -> term -> accumulator -> ctype * accumulator *) |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
534 |
fun consider_term (cdata as {hol_ctxt = {thy, ctxt, stds, fast_descrs, |
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
535 |
def_table, ...}, |
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
536 |
alpha_T, max_fresh, ...}) = |
33192 | 537 |
let |
538 |
(* typ -> ctype *) |
|
539 |
val ctype_for = fresh_ctype_for_type cdata |
|
540 |
(* ctype -> ctype *) |
|
541 |
fun pos_set_ctype_for_dom C = |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
542 |
CFun (C, S (if exists_alpha_sub_ctype C then Plus else Minus), bool_C) |
33192 | 543 |
(* typ -> accumulator -> ctype * accumulator *) |
544 |
fun do_quantifier T (gamma, cset) = |
|
545 |
let |
|
546 |
val abs_C = ctype_for (domain_type (domain_type T)) |
|
547 |
val body_C = ctype_for (range_type T) |
|
548 |
in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
549 |
(CFun (CFun (abs_C, S Minus, body_C), S Minus, body_C), |
33192 | 550 |
(gamma, cset |> add_ctype_is_right_total abs_C)) |
551 |
end |
|
552 |
fun do_equals T (gamma, cset) = |
|
553 |
let val C = ctype_for (domain_type T) in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
554 |
(CFun (C, S Minus, CFun (C, V (Unsynchronized.inc max_fresh), |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
555 |
ctype_for (nth_range_type 2 T))), |
33192 | 556 |
(gamma, cset |> add_ctype_is_right_unique C)) |
557 |
end |
|
558 |
fun do_robust_set_operation T (gamma, cset) = |
|
559 |
let |
|
560 |
val set_T = domain_type T |
|
561 |
val C1 = ctype_for set_T |
|
562 |
val C2 = ctype_for set_T |
|
563 |
val C3 = ctype_for set_T |
|
564 |
in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
565 |
(CFun (C1, S Minus, CFun (C2, S Minus, C3)), |
33192 | 566 |
(gamma, cset |> add_is_sub_ctype C1 C3 |> add_is_sub_ctype C2 C3)) |
567 |
end |
|
568 |
fun do_fragile_set_operation T (gamma, cset) = |
|
569 |
let |
|
570 |
val set_T = domain_type T |
|
571 |
val set_C = ctype_for set_T |
|
572 |
(* typ -> ctype *) |
|
573 |
fun custom_ctype_for (T as Type ("fun", [T1, T2])) = |
|
574 |
if T = set_T then set_C |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
575 |
else CFun (custom_ctype_for T1, S Minus, custom_ctype_for T2) |
33192 | 576 |
| custom_ctype_for T = ctype_for T |
577 |
in |
|
578 |
(custom_ctype_for T, (gamma, cset |> add_ctype_is_right_unique set_C)) |
|
579 |
end |
|
580 |
(* typ -> accumulator -> ctype * accumulator *) |
|
581 |
fun do_pair_constr T accum = |
|
582 |
case ctype_for (nth_range_type 2 T) of |
|
583 |
C as CPair (a_C, b_C) => |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
584 |
(CFun (a_C, S Minus, CFun (b_C, S Minus, C)), accum) |
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
585 |
| C => raise CTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [C]) |
33192 | 586 |
(* int -> typ -> accumulator -> ctype * accumulator *) |
587 |
fun do_nth_pair_sel n T = |
|
588 |
case ctype_for (domain_type T) of |
|
589 |
C as CPair (a_C, b_C) => |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
590 |
pair (CFun (C, S Minus, if n = 0 then a_C else b_C)) |
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
591 |
| C => raise CTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [C]) |
33192 | 592 |
val unsolvable = (CType ("unsolvable", []), unsolvable_accum) |
593 |
(* typ -> term -> accumulator -> ctype * accumulator *) |
|
594 |
fun do_bounded_quantifier abs_T bound_t body_t accum = |
|
595 |
let |
|
596 |
val abs_C = ctype_for abs_T |
|
597 |
val (bound_C, accum) = accum |>> push_bound abs_C |> do_term bound_t |
|
598 |
val expected_bound_C = pos_set_ctype_for_dom abs_C |
|
599 |
in |
|
600 |
accum ||> add_ctypes_equal expected_bound_C bound_C |> do_term body_t |
|
601 |
||> apfst pop_bound |
|
602 |
end |
|
603 |
(* term -> accumulator -> ctype * accumulator *) |
|
604 |
and do_term _ (_, UnsolvableCSet) = unsolvable |
|
605 |
| do_term t (accum as (gamma as {bounds, frees, consts}, cset)) = |
|
606 |
(case t of |
|
607 |
Const (x as (s, T)) => |
|
608 |
(case AList.lookup (op =) consts x of |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
609 |
SOME C => (C, accum) |
33192 | 610 |
| NONE => |
611 |
if not (could_exist_alpha_subtype alpha_T T) then |
|
612 |
(ctype_for T, accum) |
|
613 |
else case s of |
|
614 |
@{const_name all} => do_quantifier T accum |
|
615 |
| @{const_name "=="} => do_equals T accum |
|
616 |
| @{const_name All} => do_quantifier T accum |
|
617 |
| @{const_name Ex} => do_quantifier T accum |
|
618 |
| @{const_name "op ="} => do_equals T accum |
|
619 |
| @{const_name The} => (print_g "*** The"; unsolvable) |
|
620 |
| @{const_name Eps} => (print_g "*** Eps"; unsolvable) |
|
621 |
| @{const_name If} => |
|
622 |
do_robust_set_operation (range_type T) accum |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
623 |
|>> curry3 CFun bool_C (S Minus) |
33192 | 624 |
| @{const_name Pair} => do_pair_constr T accum |
625 |
| @{const_name fst} => do_nth_pair_sel 0 T accum |
|
626 |
| @{const_name snd} => do_nth_pair_sel 1 T accum |
|
627 |
| @{const_name Id} => |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
628 |
(CFun (ctype_for (domain_type T), S Minus, bool_C), accum) |
33192 | 629 |
| @{const_name insert} => |
630 |
let |
|
631 |
val set_T = domain_type (range_type T) |
|
632 |
val C1 = ctype_for (domain_type set_T) |
|
633 |
val C1' = pos_set_ctype_for_dom C1 |
|
634 |
val C2 = ctype_for set_T |
|
635 |
val C3 = ctype_for set_T |
|
636 |
in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
637 |
(CFun (C1, S Minus, CFun (C2, S Minus, C3)), |
33192 | 638 |
(gamma, cset |> add_ctype_is_right_unique C1 |
639 |
|> add_is_sub_ctype C1' C3 |
|
640 |
|> add_is_sub_ctype C2 C3)) |
|
641 |
end |
|
642 |
| @{const_name converse} => |
|
643 |
let |
|
644 |
val x = Unsynchronized.inc max_fresh |
|
645 |
(* typ -> ctype *) |
|
646 |
fun ctype_for_set T = |
|
647 |
CFun (ctype_for (domain_type T), V x, bool_C) |
|
648 |
val ab_set_C = domain_type T |> ctype_for_set |
|
649 |
val ba_set_C = range_type T |> ctype_for_set |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
650 |
in (CFun (ab_set_C, S Minus, ba_set_C), accum) end |
33192 | 651 |
| @{const_name trancl} => do_fragile_set_operation T accum |
652 |
| @{const_name rtrancl} => (print_g "*** rtrancl"; unsolvable) |
|
653 |
| @{const_name finite} => |
|
654 |
let val C1 = ctype_for (domain_type (domain_type T)) in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
655 |
(CFun (pos_set_ctype_for_dom C1, S Minus, bool_C), accum) |
33192 | 656 |
end |
657 |
| @{const_name rel_comp} => |
|
658 |
let |
|
659 |
val x = Unsynchronized.inc max_fresh |
|
660 |
(* typ -> ctype *) |
|
661 |
fun ctype_for_set T = |
|
662 |
CFun (ctype_for (domain_type T), V x, bool_C) |
|
663 |
val bc_set_C = domain_type T |> ctype_for_set |
|
664 |
val ab_set_C = domain_type (range_type T) |> ctype_for_set |
|
665 |
val ac_set_C = nth_range_type 2 T |> ctype_for_set |
|
666 |
in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
667 |
(CFun (bc_set_C, S Minus, CFun (ab_set_C, S Minus, ac_set_C)), |
33192 | 668 |
accum) |
669 |
end |
|
670 |
| @{const_name image} => |
|
671 |
let |
|
672 |
val a_C = ctype_for (domain_type (domain_type T)) |
|
673 |
val b_C = ctype_for (range_type (domain_type T)) |
|
674 |
in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
675 |
(CFun (CFun (a_C, S Minus, b_C), S Minus, |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
676 |
CFun (pos_set_ctype_for_dom a_C, S Minus, |
33192 | 677 |
pos_set_ctype_for_dom b_C)), accum) |
678 |
end |
|
679 |
| @{const_name Sigma} => |
|
680 |
let |
|
681 |
val x = Unsynchronized.inc max_fresh |
|
682 |
(* typ -> ctype *) |
|
683 |
fun ctype_for_set T = |
|
684 |
CFun (ctype_for (domain_type T), V x, bool_C) |
|
685 |
val a_set_T = domain_type T |
|
686 |
val a_C = ctype_for (domain_type a_set_T) |
|
687 |
val b_set_C = ctype_for_set (range_type (domain_type |
|
688 |
(range_type T))) |
|
689 |
val a_set_C = ctype_for_set a_set_T |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
690 |
val a_to_b_set_C = CFun (a_C, S Minus, b_set_C) |
33192 | 691 |
val ab_set_C = ctype_for_set (nth_range_type 2 T) |
692 |
in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
693 |
(CFun (a_set_C, S Minus, |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
694 |
CFun (a_to_b_set_C, S Minus, ab_set_C)), accum) |
33192 | 695 |
end |
696 |
| @{const_name Tha} => |
|
697 |
let |
|
698 |
val a_C = ctype_for (domain_type (domain_type T)) |
|
699 |
val a_set_C = pos_set_ctype_for_dom a_C |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
700 |
in (CFun (a_set_C, S Minus, a_C), accum) end |
33192 | 701 |
| @{const_name FunBox} => |
702 |
let val dom_C = ctype_for (domain_type T) in |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
703 |
(CFun (dom_C, S Minus, dom_C), accum) |
33192 | 704 |
end |
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
705 |
| _ => |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
706 |
if s = @{const_name minus_class.minus} andalso |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
707 |
is_set_type (domain_type T) then |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
708 |
let |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
709 |
val set_T = domain_type T |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
710 |
val left_set_C = ctype_for set_T |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
711 |
val right_set_C = ctype_for set_T |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
712 |
in |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
713 |
(CFun (left_set_C, S Minus, |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
714 |
CFun (right_set_C, S Minus, left_set_C)), |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
715 |
(gamma, cset |> add_ctype_is_right_unique right_set_C |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
716 |
|> add_is_sub_ctype right_set_C left_set_C)) |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
717 |
end |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
718 |
else if s = @{const_name ord_class.less_eq} andalso |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
719 |
is_set_type (domain_type T) then |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
720 |
do_fragile_set_operation T accum |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
721 |
else if (s = @{const_name semilattice_inf_class.inf} orelse |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
722 |
s = @{const_name semilattice_sup_class.sup}) andalso |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
723 |
is_set_type (domain_type T) then |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
724 |
do_robust_set_operation T accum |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
725 |
else if is_sel s then |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
726 |
if constr_name_for_sel_like s = @{const_name FunBox} then |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
727 |
let val dom_C = ctype_for (domain_type T) in |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
728 |
(CFun (dom_C, S Minus, dom_C), accum) |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
729 |
end |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
730 |
else |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
731 |
(ctype_for_sel cdata x, accum) |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
732 |
else if is_constr thy stds x then |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
733 |
(ctype_for_constr cdata x, accum) |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
734 |
else if is_built_in_const thy stds fast_descrs x then |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
735 |
case def_of_const thy def_table x of |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
736 |
SOME t' => do_term t' accum |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
737 |
| NONE => (print_g ("*** built-in " ^ s); unsolvable) |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
738 |
else |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
739 |
let val C = ctype_for T in |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
740 |
(C, ({bounds = bounds, frees = frees, |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
741 |
consts = (x, C) :: consts}, cset)) |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
742 |
end) |
33192 | 743 |
| Free (x as (_, T)) => |
744 |
(case AList.lookup (op =) frees x of |
|
745 |
SOME C => (C, accum) |
|
746 |
| NONE => |
|
747 |
let val C = ctype_for T in |
|
748 |
(C, ({bounds = bounds, frees = (x, C) :: frees, |
|
749 |
consts = consts}, cset)) |
|
750 |
end) |
|
751 |
| Var _ => (print_g "*** Var"; unsolvable) |
|
752 |
| Bound j => (nth bounds j, accum) |
|
753 |
| Abs (_, T, @{const False}) => (ctype_for (T --> bool_T), accum) |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
754 |
| Abs (_, T, t') => |
34998 | 755 |
((case t' of |
756 |
t1' $ Bound 0 => |
|
757 |
if not (loose_bvar1 (t1', 0)) then |
|
758 |
do_term (incr_boundvars ~1 t1') accum |
|
759 |
else |
|
760 |
raise SAME () |
|
761 |
| _ => raise SAME ()) |
|
762 |
handle SAME () => |
|
763 |
let |
|
764 |
val C = ctype_for T |
|
765 |
val (C', accum) = do_term t' (accum |>> push_bound C) |
|
766 |
in (CFun (C, S Minus, C'), accum |>> pop_bound) end) |
|
33192 | 767 |
| Const (@{const_name All}, _) |
768 |
$ Abs (_, T', @{const "op -->"} $ (t1 $ Bound 0) $ t2) => |
|
769 |
do_bounded_quantifier T' t1 t2 accum |
|
770 |
| Const (@{const_name Ex}, _) |
|
771 |
$ Abs (_, T', @{const "op &"} $ (t1 $ Bound 0) $ t2) => |
|
772 |
do_bounded_quantifier T' t1 t2 accum |
|
773 |
| Const (@{const_name Let}, _) $ t1 $ t2 => |
|
774 |
do_term (betapply (t2, t1)) accum |
|
775 |
| t1 $ t2 => |
|
776 |
let |
|
777 |
val (C1, accum) = do_term t1 accum |
|
778 |
val (C2, accum) = do_term t2 accum |
|
779 |
in |
|
780 |
case accum of |
|
781 |
(_, UnsolvableCSet) => unsolvable |
|
782 |
| _ => case C1 of |
|
783 |
CFun (C11, _, C12) => |
|
784 |
(C12, accum ||> add_is_sub_ctype C2 C11) |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
785 |
| _ => raise CTYPE ("Nitpick_Mono.consider_term.do_term \ |
33192 | 786 |
\(op $)", [C1]) |
787 |
end) |
|
788 |
|> tap (fn (C, _) => |
|
789 |
print_g (" \<Gamma> \<turnstile> " ^ |
|
790 |
Syntax.string_of_term ctxt t ^ " : " ^ |
|
791 |
string_for_ctype C)) |
|
792 |
in do_term end |
|
793 |
||
794 |
(* cdata -> sign -> term -> accumulator -> accumulator *) |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
795 |
fun consider_general_formula (cdata as {hol_ctxt = {ctxt, ...}, ...}) = |
33192 | 796 |
let |
797 |
(* typ -> ctype *) |
|
798 |
val ctype_for = fresh_ctype_for_type cdata |
|
33732
385381514eed
added constraint for Eq^- in Nitpick's implementation of the monotonicity calculus
blanchet
parents:
33580
diff
changeset
|
799 |
(* term -> accumulator -> ctype * accumulator *) |
385381514eed
added constraint for Eq^- in Nitpick's implementation of the monotonicity calculus
blanchet
parents:
33580
diff
changeset
|
800 |
val do_term = consider_term cdata |
33192 | 801 |
(* sign -> term -> accumulator -> accumulator *) |
802 |
fun do_formula _ _ (_, UnsolvableCSet) = unsolvable_accum |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
803 |
| do_formula sn t (accum as (gamma, cset)) = |
33192 | 804 |
let |
805 |
(* term -> accumulator -> accumulator *) |
|
806 |
val do_co_formula = do_formula sn |
|
807 |
val do_contra_formula = do_formula (negate sn) |
|
808 |
(* string -> typ -> term -> accumulator *) |
|
809 |
fun do_quantifier quant_s abs_T body_t = |
|
810 |
let |
|
811 |
val abs_C = ctype_for abs_T |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
812 |
val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex})) |
33192 | 813 |
val cset = cset |> side_cond ? add_ctype_is_right_total abs_C |
814 |
in |
|
815 |
(gamma |> push_bound abs_C, cset) |> do_co_formula body_t |
|
816 |
|>> pop_bound |
|
817 |
end |
|
818 |
(* typ -> term -> accumulator *) |
|
819 |
fun do_bounded_quantifier abs_T body_t = |
|
820 |
accum |>> push_bound (ctype_for abs_T) |> do_co_formula body_t |
|
821 |
|>> pop_bound |
|
822 |
(* term -> term -> accumulator *) |
|
823 |
fun do_equals t1 t2 = |
|
824 |
case sn of |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
825 |
Plus => do_term t accum |> snd |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
826 |
| Minus => let |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
827 |
val (C1, accum) = do_term t1 accum |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
828 |
val (C2, accum) = do_term t2 accum |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
829 |
in accum ||> add_ctypes_equal C1 C2 end |
33192 | 830 |
in |
831 |
case t of |
|
832 |
Const (s0 as @{const_name all}, _) $ Abs (_, T1, t1) => |
|
833 |
do_quantifier s0 T1 t1 |
|
834 |
| Const (@{const_name "=="}, _) $ t1 $ t2 => do_equals t1 t2 |
|
835 |
| @{const "==>"} $ t1 $ t2 => |
|
836 |
accum |> do_contra_formula t1 |> do_co_formula t2 |
|
837 |
| @{const Trueprop} $ t1 => do_co_formula t1 accum |
|
838 |
| @{const Not} $ t1 => do_contra_formula t1 accum |
|
839 |
| Const (@{const_name All}, _) |
|
840 |
$ Abs (_, T1, t1 as @{const "op -->"} $ (_ $ Bound 0) $ _) => |
|
841 |
do_bounded_quantifier T1 t1 |
|
842 |
| Const (s0 as @{const_name All}, _) $ Abs (_, T1, t1) => |
|
843 |
do_quantifier s0 T1 t1 |
|
844 |
| Const (@{const_name Ex}, _) |
|
845 |
$ Abs (_, T1, t1 as @{const "op &"} $ (_ $ Bound 0) $ _) => |
|
846 |
do_bounded_quantifier T1 t1 |
|
847 |
| Const (s0 as @{const_name Ex}, _) $ Abs (_, T1, t1) => |
|
848 |
do_quantifier s0 T1 t1 |
|
849 |
| Const (@{const_name "op ="}, _) $ t1 $ t2 => do_equals t1 t2 |
|
850 |
| @{const "op &"} $ t1 $ t2 => |
|
851 |
accum |> do_co_formula t1 |> do_co_formula t2 |
|
852 |
| @{const "op |"} $ t1 $ t2 => |
|
853 |
accum |> do_co_formula t1 |> do_co_formula t2 |
|
854 |
| @{const "op -->"} $ t1 $ t2 => |
|
855 |
accum |> do_contra_formula t1 |> do_co_formula t2 |
|
856 |
| Const (@{const_name If}, _) $ t1 $ t2 $ t3 => |
|
33732
385381514eed
added constraint for Eq^- in Nitpick's implementation of the monotonicity calculus
blanchet
parents:
33580
diff
changeset
|
857 |
accum |> do_term t1 |> snd |> fold do_co_formula [t2, t3] |
33192 | 858 |
| Const (@{const_name Let}, _) $ t1 $ t2 => |
859 |
do_co_formula (betapply (t2, t1)) accum |
|
33732
385381514eed
added constraint for Eq^- in Nitpick's implementation of the monotonicity calculus
blanchet
parents:
33580
diff
changeset
|
860 |
| _ => do_term t accum |> snd |
33192 | 861 |
end |
862 |
|> tap (fn _ => print_g ("\<Gamma> \<turnstile> " ^ |
|
863 |
Syntax.string_of_term ctxt t ^ |
|
864 |
" : o\<^sup>" ^ string_for_sign sn)) |
|
865 |
in do_formula end |
|
866 |
||
867 |
(* The harmless axiom optimization below is somewhat too aggressive in the face |
|
868 |
of (rather peculiar) user-defined axioms. *) |
|
869 |
val harmless_consts = |
|
870 |
[@{const_name ord_class.less}, @{const_name ord_class.less_eq}] |
|
871 |
val bounteous_consts = [@{const_name bisim}] |
|
872 |
||
873 |
(* term -> bool *) |
|
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
874 |
fun is_harmless_axiom ({thy, stds, fast_descrs, ...} : hol_context) t = |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
875 |
Term.add_consts t [] |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
876 |
|> filter_out (is_built_in_const thy stds fast_descrs) |
33192 | 877 |
|> (forall (member (op =) harmless_consts o original_name o fst) |
878 |
orf exists (member (op =) bounteous_consts o fst)) |
|
879 |
||
880 |
(* cdata -> sign -> term -> accumulator -> accumulator *) |
|
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
881 |
fun consider_nondefinitional_axiom (cdata as {hol_ctxt, ...}) sn t = |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
882 |
not (is_harmless_axiom hol_ctxt t) ? consider_general_formula cdata sn t |
33192 | 883 |
|
884 |
(* cdata -> term -> accumulator -> accumulator *) |
|
35070
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents:
34998
diff
changeset
|
885 |
fun consider_definitional_axiom (cdata as {hol_ctxt as {thy, ...}, ...}) t = |
33192 | 886 |
if not (is_constr_pattern_formula thy t) then |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
887 |
consider_nondefinitional_axiom cdata Plus t |
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
888 |
else if is_harmless_axiom hol_ctxt t then |
33192 | 889 |
I |
890 |
else |
|
891 |
let |
|
33732
385381514eed
added constraint for Eq^- in Nitpick's implementation of the monotonicity calculus
blanchet
parents:
33580
diff
changeset
|
892 |
(* term -> accumulator -> ctype * accumulator *) |
33192 | 893 |
val do_term = consider_term cdata |
894 |
(* typ -> term -> accumulator -> accumulator *) |
|
895 |
fun do_all abs_T body_t accum = |
|
896 |
let val abs_C = fresh_ctype_for_type cdata abs_T in |
|
897 |
accum |>> push_bound abs_C |> do_formula body_t |>> pop_bound |
|
898 |
end |
|
899 |
(* term -> term -> accumulator -> accumulator *) |
|
900 |
and do_implies t1 t2 = do_term t1 #> snd #> do_formula t2 |
|
901 |
and do_equals t1 t2 accum = |
|
902 |
let |
|
903 |
val (C1, accum) = do_term t1 accum |
|
904 |
val (C2, accum) = do_term t2 accum |
|
905 |
in accum ||> add_ctypes_equal C1 C2 end |
|
906 |
(* term -> accumulator -> accumulator *) |
|
907 |
and do_formula _ (_, UnsolvableCSet) = unsolvable_accum |
|
908 |
| do_formula t accum = |
|
909 |
case t of |
|
910 |
Const (@{const_name all}, _) $ Abs (_, T1, t1) => do_all T1 t1 accum |
|
911 |
| @{const Trueprop} $ t1 => do_formula t1 accum |
|
912 |
| Const (@{const_name "=="}, _) $ t1 $ t2 => do_equals t1 t2 accum |
|
913 |
| @{const "==>"} $ t1 $ t2 => do_implies t1 t2 accum |
|
914 |
| @{const Pure.conjunction} $ t1 $ t2 => |
|
915 |
accum |> do_formula t1 |> do_formula t2 |
|
916 |
| Const (@{const_name All}, _) $ Abs (_, T1, t1) => do_all T1 t1 accum |
|
917 |
| Const (@{const_name "op ="}, _) $ t1 $ t2 => do_equals t1 t2 accum |
|
918 |
| @{const "op &"} $ t1 $ t2 => accum |> do_formula t1 |> do_formula t2 |
|
919 |
| @{const "op -->"} $ t1 $ t2 => do_implies t1 t2 accum |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
920 |
| _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\ |
33192 | 921 |
\do_formula", [t]) |
922 |
in do_formula t end |
|
923 |
||
924 |
(* Proof.context -> literal list -> term -> ctype -> string *) |
|
925 |
fun string_for_ctype_of_term ctxt lits t C = |
|
926 |
Syntax.string_of_term ctxt t ^ " : " ^ |
|
927 |
string_for_ctype (instantiate_ctype lits C) |
|
928 |
||
929 |
(* theory -> literal list -> ctype_context -> unit *) |
|
930 |
fun print_ctype_context ctxt lits ({frees, consts, ...} : ctype_context) = |
|
931 |
map (fn (x, C) => string_for_ctype_of_term ctxt lits (Free x) C) frees @ |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
932 |
map (fn (x, C) => string_for_ctype_of_term ctxt lits (Const x) C) consts |
33192 | 933 |
|> cat_lines |> print_g |
934 |
||
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
935 |
(* hol_context -> bool -> typ -> sign -> term list -> term list -> term |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
936 |
-> bool *) |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
937 |
fun formulas_monotonic (hol_ctxt as {ctxt, ...}) binarize alpha_T sn def_ts |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
938 |
nondef_ts core_t = |
33192 | 939 |
let |
940 |
val _ = print_g ("****** " ^ string_for_ctype CAlpha ^ " is " ^ |
|
941 |
Syntax.string_of_typ ctxt alpha_T) |
|
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
942 |
val cdata as {max_fresh, ...} = initial_cdata hol_ctxt binarize alpha_T |
33192 | 943 |
val (gamma, cset) = |
944 |
(initial_gamma, slack) |
|
945 |
|> fold (consider_definitional_axiom cdata) def_ts |
|
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
946 |
|> fold (consider_nondefinitional_axiom cdata Plus) nondef_ts |
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
947 |
|> consider_general_formula cdata sn core_t |
33192 | 948 |
in |
949 |
case solve (!max_fresh) cset of |
|
950 |
SOME lits => (print_ctype_context ctxt lits gamma; true) |
|
951 |
| _ => false |
|
952 |
end |
|
953 |
handle CTYPE (loc, Cs) => raise BAD (loc, commas (map string_for_ctype Cs)) |
|
954 |
||
955 |
end; |