| author | wenzelm | 
| Sat, 16 Aug 2014 14:32:26 +0200 | |
| changeset 57952 | 1a9a6dfc255f | 
| parent 56245 | 84fc7dfa3cd4 | 
| child 57996 | ca917ea6969c | 
| permissions | -rw-r--r-- | 
| 33982 | 1  | 
(* Title: HOL/Tools/Nitpick/nitpick.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: 
34938 
diff
changeset
 | 
3  | 
Copyright 2008, 2009, 2010  | 
| 33192 | 4  | 
|
5  | 
Finite model generation for HOL formulas using Kodkod.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature NITPICK =  | 
|
9  | 
sig  | 
|
| 
35711
 
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
 
blanchet 
parents: 
35696 
diff
changeset
 | 
10  | 
type term_postprocessor = Nitpick_Model.term_postprocessor  | 
| 43022 | 11  | 
|
| 47560 | 12  | 
datatype mode = Auto_Try | Try | Normal | TPTP  | 
| 43022 | 13  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
14  | 
type params =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
15  | 
    {cards_assigns: (typ option * int list) list,
 | 
| 55889 | 16  | 
maxes_assigns: ((string * typ) option * int list) list,  | 
17  | 
iters_assigns: ((string * typ) option * int list) list,  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
18  | 
bitss: int list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
19  | 
bisim_depths: int list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
20  | 
boxes: (typ option * bool option) list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
21  | 
finitizes: (typ option * bool option) list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
22  | 
monos: (typ option * bool option) list,  | 
| 55889 | 23  | 
wfs: ((string * typ) option * bool option) list,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
24  | 
sat_solver: string,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
25  | 
blocking: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
26  | 
falsify: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
27  | 
debug: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
28  | 
verbose: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
29  | 
overlord: bool,  | 
| 53802 | 30  | 
spy: bool,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
31  | 
user_axioms: bool option,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
32  | 
assms: bool,  | 
| 38209 | 33  | 
whacks: term list,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
34  | 
merge_type_vars: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
35  | 
binary_ints: bool option,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
36  | 
destroy_constrs: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
37  | 
specialize: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
38  | 
star_linear_preds: bool,  | 
| 41856 | 39  | 
total_consts: bool option,  | 
| 41876 | 40  | 
needs: term list option,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
41  | 
peephole_optim: bool,  | 
| 38124 | 42  | 
datatype_sym_break: int,  | 
43  | 
kodkod_sym_break: int,  | 
|
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
44  | 
timeout: Time.time,  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
45  | 
tac_timeout: Time.time,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
46  | 
max_threads: int,  | 
| 55889 | 47  | 
show_types: bool,  | 
| 
41993
 
bd6296de1432
reintroduced "show_skolems" option -- useful when too many Skolems are displayed
 
blanchet 
parents: 
41992 
diff
changeset
 | 
48  | 
show_skolems: bool,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
49  | 
show_consts: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
50  | 
evals: term list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
51  | 
formats: (term option * int list) list,  | 
| 
37260
 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 
blanchet 
parents: 
37257 
diff
changeset
 | 
52  | 
atomss: (typ option * string list) list,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
53  | 
max_potential: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
54  | 
max_genuine: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
55  | 
check_potential: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
56  | 
check_genuine: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
57  | 
batch_size: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
58  | 
expect: string}  | 
| 33192 | 59  | 
|
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
60  | 
val genuineN : string  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
61  | 
val quasi_genuineN : string  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
62  | 
val potentialN : string  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
63  | 
val noneN : string  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
64  | 
val unknownN : string  | 
| 33192 | 65  | 
val register_frac_type : string -> (string * string) list -> theory -> theory  | 
66  | 
val unregister_frac_type : string -> theory -> theory  | 
|
| 55889 | 67  | 
val register_codatatype : typ -> string -> (string * typ) list -> theory ->  | 
68  | 
theory  | 
|
| 33192 | 69  | 
val unregister_codatatype : typ -> theory -> theory  | 
| 
35711
 
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
 
blanchet 
parents: 
35696 
diff
changeset
 | 
70  | 
val register_term_postprocessor :  | 
| 
 
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
 
blanchet 
parents: 
35696 
diff
changeset
 | 
71  | 
typ -> term_postprocessor -> theory -> theory  | 
| 
 
548d3f16404b
added term postprocessor to Nitpick, to provide custom syntax for typedefs
 
blanchet 
parents: 
35696 
diff
changeset
 | 
72  | 
val unregister_term_postprocessor : typ -> theory -> theory  | 
| 33192 | 73  | 
val pick_nits_in_term :  | 
| 43022 | 74  | 
Proof.state -> params -> mode -> int -> int -> int -> (term * term) list  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
75  | 
-> term list -> term list -> term -> string * Proof.state  | 
| 33192 | 76  | 
val pick_nits_in_subgoal :  | 
| 43022 | 77  | 
Proof.state -> params -> mode -> int -> int -> string * Proof.state  | 
| 33192 | 78  | 
end;  | 
79  | 
||
80  | 
structure Nitpick : NITPICK =  | 
|
81  | 
struct  | 
|
82  | 
||
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
83  | 
open Nitpick_Util  | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
84  | 
open Nitpick_HOL  | 
| 
35070
 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 
blanchet 
parents: 
34998 
diff
changeset
 | 
85  | 
open Nitpick_Preproc  | 
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
86  | 
open Nitpick_Mono  | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
87  | 
open Nitpick_Scope  | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
88  | 
open Nitpick_Peephole  | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
89  | 
open Nitpick_Rep  | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
90  | 
open Nitpick_Nut  | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
91  | 
open Nitpick_Kodkod  | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
92  | 
open Nitpick_Model  | 
| 33192 | 93  | 
|
| 34126 | 94  | 
structure KK = Kodkod  | 
95  | 
||
| 47560 | 96  | 
datatype mode = Auto_Try | Try | Normal | TPTP  | 
97  | 
||
98  | 
fun is_mode_nt mode = (mode = Normal orelse mode = TPTP)  | 
|
| 43022 | 99  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
100  | 
type params =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
101  | 
  {cards_assigns: (typ option * int list) list,
 | 
| 55889 | 102  | 
maxes_assigns: ((string * typ) option * int list) list,  | 
103  | 
iters_assigns: ((string * typ) option * int list) list,  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
104  | 
bitss: int list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
105  | 
bisim_depths: int list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
106  | 
boxes: (typ option * bool option) list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
107  | 
finitizes: (typ option * bool option) list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
108  | 
monos: (typ option * bool option) list,  | 
| 55889 | 109  | 
wfs: ((string * typ) option * bool option) list,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
110  | 
sat_solver: string,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
111  | 
blocking: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
112  | 
falsify: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
113  | 
debug: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
114  | 
verbose: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
115  | 
overlord: bool,  | 
| 53802 | 116  | 
spy: bool,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
117  | 
user_axioms: bool option,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
118  | 
assms: bool,  | 
| 38209 | 119  | 
whacks: term list,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
120  | 
merge_type_vars: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
121  | 
binary_ints: bool option,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
122  | 
destroy_constrs: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
123  | 
specialize: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
124  | 
star_linear_preds: bool,  | 
| 41856 | 125  | 
total_consts: bool option,  | 
| 41876 | 126  | 
needs: term list option,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
127  | 
peephole_optim: bool,  | 
| 38124 | 128  | 
datatype_sym_break: int,  | 
129  | 
kodkod_sym_break: int,  | 
|
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
130  | 
timeout: Time.time,  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
131  | 
tac_timeout: Time.time,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
132  | 
max_threads: int,  | 
| 55889 | 133  | 
show_types: bool,  | 
| 
41993
 
bd6296de1432
reintroduced "show_skolems" option -- useful when too many Skolems are displayed
 
blanchet 
parents: 
41992 
diff
changeset
 | 
134  | 
show_skolems: bool,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
135  | 
show_consts: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
136  | 
evals: term list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
137  | 
formats: (term option * int list) list,  | 
| 
37260
 
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
 
blanchet 
parents: 
37257 
diff
changeset
 | 
138  | 
atomss: (typ option * string list) list,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
139  | 
max_potential: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
140  | 
max_genuine: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
141  | 
check_potential: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
142  | 
check_genuine: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
143  | 
batch_size: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
144  | 
expect: string}  | 
| 33192 | 145  | 
|
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
146  | 
val genuineN = "genuine"  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
147  | 
val quasi_genuineN = "quasi_genuine"  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
148  | 
val potentialN = "potential"  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
149  | 
val noneN = "none"  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
150  | 
val unknownN = "unknown"  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
151  | 
|
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
152  | 
(* TODO: eliminate these historical aliases *)  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
153  | 
val register_frac_type = Nitpick_HOL.register_frac_type_global  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
154  | 
val unregister_frac_type = Nitpick_HOL.unregister_frac_type_global  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
155  | 
val register_codatatype = Nitpick_HOL.register_codatatype_global  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
156  | 
val unregister_codatatype = Nitpick_HOL.unregister_codatatype_global  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
157  | 
val register_term_postprocessor =  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
158  | 
Nitpick_Model.register_term_postprocessor_global  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
159  | 
val unregister_term_postprocessor =  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
160  | 
Nitpick_Model.unregister_term_postprocessor_global  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
161  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
162  | 
type problem_extension =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
163  | 
  {free_names: nut list,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
164  | 
sel_names: nut list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
165  | 
nonsel_names: nut list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
166  | 
rel_table: nut NameTable.table,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
167  | 
unsound: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36389 
diff
changeset
 | 
168  | 
scope: scope}  | 
| 
39316
 
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
 
blanchet 
parents: 
38857 
diff
changeset
 | 
169  | 
|
| 34126 | 170  | 
type rich_problem = KK.problem * problem_extension  | 
| 33192 | 171  | 
|
172  | 
fun pretties_for_formulas _ _ [] = []  | 
|
173  | 
| pretties_for_formulas ctxt s ts =  | 
|
174  | 
[Pretty.str (s ^ plural_s_for_list ts ^ ":"),  | 
|
175  | 
Pretty.indent indent_size (Pretty.chunks  | 
|
176  | 
(map2 (fn j => fn t =>  | 
|
| 
34121
 
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
 
blanchet 
parents: 
34039 
diff
changeset
 | 
177  | 
Pretty.block [t |> shorten_names_in_term  | 
| 33192 | 178  | 
|> Syntax.pretty_term ctxt,  | 
179  | 
Pretty.str (if j = 1 then "." else ";")])  | 
|
180  | 
(length ts downto 1) ts))]  | 
|
181  | 
||
| 49024 | 182  | 
val isabelle_wrong_message =  | 
183  | 
"Something appears to be wrong with your Isabelle installation."  | 
|
| 55889 | 184  | 
val java_not_found_message =  | 
| 49024 | 185  | 
"Java could not be launched. " ^ isabelle_wrong_message  | 
| 55889 | 186  | 
val java_too_old_message =  | 
| 49024 | 187  | 
"The Java version is too old. " ^ isabelle_wrong_message  | 
| 55889 | 188  | 
val kodkodi_not_installed_message =  | 
| 50830 | 189  | 
"Nitpick requires the external Java program Kodkodi."  | 
| 55889 | 190  | 
val kodkodi_too_old_message = "The installed Kodkodi version is too old."  | 
191  | 
||
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
192  | 
val max_unsound_delay_ms = 200  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
193  | 
val max_unsound_delay_percent = 2  | 
| 33192 | 194  | 
|
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
195  | 
fun unsound_delay_for_timeout timeout =  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
196  | 
Int.max (0, Int.min (max_unsound_delay_ms,  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
197  | 
Time.toMilliseconds timeout  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
198  | 
* max_unsound_delay_percent div 100))  | 
| 33192 | 199  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
200  | 
fun none_true assigns = forall (not_equal (SOME true) o snd) assigns  | 
| 33192 | 201  | 
|
| 
41048
 
d5ebe94248ad
added a hint when the user obviously just forgot a colon after the lemma's name
 
blanchet 
parents: 
41007 
diff
changeset
 | 
202  | 
fun has_lonely_bool_var (@{const Pure.conjunction}
 | 
| 
 
d5ebe94248ad
added a hint when the user obviously just forgot a colon after the lemma's name
 
blanchet 
parents: 
41007 
diff
changeset
 | 
203  | 
                         $ (@{const Trueprop} $ Free _) $ _) = true
 | 
| 
 
d5ebe94248ad
added a hint when the user obviously just forgot a colon after the lemma's name
 
blanchet 
parents: 
41007 
diff
changeset
 | 
204  | 
| has_lonely_bool_var _ = false  | 
| 
 
d5ebe94248ad
added a hint when the user obviously just forgot a colon after the lemma's name
 
blanchet 
parents: 
41007 
diff
changeset
 | 
205  | 
|
| 
34038
 
a2736debeabd
make Nitpick output the message "Hint: Maybe you forgot a type constraint?" only for syntactic classes
 
blanchet 
parents: 
33982 
diff
changeset
 | 
206  | 
val syntactic_sorts =  | 
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38517 
diff
changeset
 | 
207  | 
  @{sort "{default,zero,one,plus,minus,uminus,times,inverse,abs,sgn,ord,equal}"} @
 | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46243 
diff
changeset
 | 
208  | 
  @{sort numeral}
 | 
| 55889 | 209  | 
|
| 
34038
 
a2736debeabd
make Nitpick output the message "Hint: Maybe you forgot a type constraint?" only for syntactic classes
 
blanchet 
parents: 
33982 
diff
changeset
 | 
210  | 
fun has_tfree_syntactic_sort (TFree (_, S as _ :: _)) =  | 
| 
 
a2736debeabd
make Nitpick output the message "Hint: Maybe you forgot a type constraint?" only for syntactic classes
 
blanchet 
parents: 
33982 
diff
changeset
 | 
211  | 
subset (op =) (S, syntactic_sorts)  | 
| 
 
a2736debeabd
make Nitpick output the message "Hint: Maybe you forgot a type constraint?" only for syntactic classes
 
blanchet 
parents: 
33982 
diff
changeset
 | 
212  | 
| has_tfree_syntactic_sort _ = false  | 
| 
 
a2736debeabd
make Nitpick output the message "Hint: Maybe you forgot a type constraint?" only for syntactic classes
 
blanchet 
parents: 
33982 
diff
changeset
 | 
213  | 
val has_syntactic_sorts = exists_type (exists_subtype has_tfree_syntactic_sort)  | 
| 33192 | 214  | 
|
| 
33568
 
532b915afa14
don't run Nitpick at all if Kodkodi is not installed (as indicated by the $KODKODI variable)
 
blanchet 
parents: 
33566 
diff
changeset
 | 
215  | 
fun plazy f = Pretty.blk (0, pstrs (f ()))  | 
| 
 
532b915afa14
don't run Nitpick at all if Kodkodi is not installed (as indicated by the $KODKODI variable)
 
blanchet 
parents: 
33566 
diff
changeset
 | 
216  | 
|
| 43022 | 217  | 
fun pick_them_nits_in_term deadline state (params : params) mode i n step  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
218  | 
subst def_assm_ts nondef_assm_ts orig_t =  | 
| 33192 | 219  | 
let  | 
220  | 
val timer = Timer.startRealTimer ()  | 
|
| 
34935
 
cb011ba38950
removed the Nitpick code that loaded the "Nitpick" theory explicitly if it's not already loaded, because this didn't work properly and is of doubtful value
 
blanchet 
parents: 
34126 
diff
changeset
 | 
221  | 
val thy = Proof.theory_of state  | 
| 
 
cb011ba38950
removed the Nitpick code that loaded the "Nitpick" theory explicitly if it's not already loaded, because this didn't work properly and is of doubtful value
 
blanchet 
parents: 
34126 
diff
changeset
 | 
222  | 
val ctxt = Proof.context_of state  | 
| 
35386
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
223  | 
(* FIXME: reintroduce code before new release:  | 
| 
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
224  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37213 
diff
changeset
 | 
225  | 
val nitpick_thy = Thy_Info.get_theory "Nitpick"  | 
| 
34936
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
226  | 
val _ = Theory.subthy (nitpick_thy, thy) orelse  | 
| 
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
227  | 
error "You must import the theory \"Nitpick\" to use Nitpick"  | 
| 
34982
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
228  | 
*)  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
229  | 
    val {cards_assigns, maxes_assigns, iters_assigns, bitss, bisim_depths,
 | 
| 55888 | 230  | 
boxes, finitizes, monos, wfs, sat_solver, falsify, debug, verbose,  | 
231  | 
overlord, spy, user_axioms, assms, whacks, merge_type_vars,  | 
|
| 38209 | 232  | 
binary_ints, destroy_constrs, specialize, star_linear_preds,  | 
| 41875 | 233  | 
total_consts, needs, peephole_optim, datatype_sym_break,  | 
| 55889 | 234  | 
kodkod_sym_break, tac_timeout, max_threads, show_types, show_skolems,  | 
235  | 
show_consts, evals, formats, atomss, max_potential, max_genuine,  | 
|
236  | 
check_potential, check_genuine, batch_size, ...} = params  | 
|
| 33192 | 237  | 
val state_ref = Unsynchronized.ref state  | 
| 46086 | 238  | 
fun pprint print =  | 
| 43022 | 239  | 
if mode = Auto_Try then  | 
| 33192 | 240  | 
Unsynchronized.change state_ref o Proof.goal_message o K  | 
| 
52643
 
34c29356930e
more explicit Markup.information for messages produced by "auto" tools;
 
wenzelm 
parents: 
52202 
diff
changeset
 | 
241  | 
o Pretty.mark Markup.information  | 
| 33192 | 242  | 
else  | 
| 46086 | 243  | 
print o Pretty.string_of  | 
| 47559 | 244  | 
val pprint_a = pprint Output.urgent_message  | 
| 47560 | 245  | 
fun pprint_nt f = () |> is_mode_nt mode ? pprint Output.urgent_message o f  | 
| 46086 | 246  | 
fun pprint_v f = () |> verbose ? pprint Output.urgent_message o f  | 
247  | 
fun pprint_d f = () |> debug ? pprint tracing o f  | 
|
248  | 
val print = pprint Output.urgent_message o curry Pretty.blk 0 o pstrs  | 
|
| 47560 | 249  | 
fun print_t f = () |> mode = TPTP ? print o f  | 
| 39345 | 250  | 
(*  | 
| 46086 | 251  | 
val print_g = pprint tracing o Pretty.str  | 
| 39345 | 252  | 
*)  | 
| 47560 | 253  | 
val print_nt = pprint_nt o K o plazy  | 
| 
33568
 
532b915afa14
don't run Nitpick at all if Kodkodi is not installed (as indicated by the $KODKODI variable)
 
blanchet 
parents: 
33566 
diff
changeset
 | 
254  | 
val print_v = pprint_v o K o plazy  | 
| 33192 | 255  | 
|
256  | 
fun check_deadline () =  | 
|
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
257  | 
if Time.compare (Time.now (), deadline) <> LESS then  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
258  | 
raise TimeLimit.TimeOut  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
259  | 
else  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
260  | 
()  | 
| 33192 | 261  | 
|
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
262  | 
val (def_assm_ts, nondef_assm_ts) =  | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
263  | 
if assms orelse mode <> Normal then (def_assm_ts, nondef_assm_ts)  | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
264  | 
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: 
34938 
diff
changeset
 | 
265  | 
val _ =  | 
| 
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
266  | 
if step = 0 then  | 
| 47560 | 267  | 
print_nt (fn () => "Nitpicking formula...")  | 
| 
34982
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
268  | 
else  | 
| 47560 | 269  | 
pprint_nt (fn () => Pretty.chunks (  | 
| 
34982
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
270  | 
            pretties_for_formulas ctxt ("Nitpicking " ^
 | 
| 
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
271  | 
(if i <> 1 orelse n <> 1 then  | 
| 
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
272  | 
"subgoal " ^ string_of_int i ^ " of " ^ string_of_int n  | 
| 
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
273  | 
else  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
274  | 
"goal")) [Logic.list_implies (nondef_assm_ts, orig_t)]))  | 
| 
53815
 
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
 
blanchet 
parents: 
53802 
diff
changeset
 | 
275  | 
    val _ = spying spy (fn () => (state, i, "starting " ^ @{make_string} mode ^ " mode"))
 | 
| 
41278
 
8e1cde88aae6
added a timestamp to Nitpick in verbose mode for debugging purposes;
 
blanchet 
parents: 
41052 
diff
changeset
 | 
276  | 
val _ = print_v (enclose "Timestamp: " "." o Date.fmt "%H:%M:%S"  | 
| 
 
8e1cde88aae6
added a timestamp to Nitpick in verbose mode for debugging purposes;
 
blanchet 
parents: 
41052 
diff
changeset
 | 
277  | 
o Date.fromTimeLocal o Time.now)  | 
| 33192 | 278  | 
    val neg_t = if falsify then Logic.mk_implies (orig_t, @{prop False})
 | 
279  | 
else orig_t  | 
|
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
280  | 
val conj_ts = neg_t :: def_assm_ts @ nondef_assm_ts @ evals @ these needs  | 
| 
38169
 
b51784515471
optimize local "def"s by treating them as definitions
 
blanchet 
parents: 
38126 
diff
changeset
 | 
281  | 
val tfree_table =  | 
| 
41869
 
9e367f1c9570
robust handling of types occurring in "eval" and "preconstr" options but not in the goal
 
blanchet 
parents: 
41868 
diff
changeset
 | 
282  | 
if merge_type_vars then merged_type_var_table_for_terms thy conj_ts  | 
| 
 
9e367f1c9570
robust handling of types occurring in "eval" and "preconstr" options but not in the goal
 
blanchet 
parents: 
41868 
diff
changeset
 | 
283  | 
else []  | 
| 
 
9e367f1c9570
robust handling of types occurring in "eval" and "preconstr" options but not in the goal
 
blanchet 
parents: 
41868 
diff
changeset
 | 
284  | 
val merge_tfrees = merge_type_vars_in_term thy merge_type_vars tfree_table  | 
| 
 
9e367f1c9570
robust handling of types occurring in "eval" and "preconstr" options but not in the goal
 
blanchet 
parents: 
41868 
diff
changeset
 | 
285  | 
val neg_t = neg_t |> merge_tfrees  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
286  | 
val def_assm_ts = def_assm_ts |> map merge_tfrees  | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
287  | 
val nondef_assm_ts = nondef_assm_ts |> map merge_tfrees  | 
| 
41869
 
9e367f1c9570
robust handling of types occurring in "eval" and "preconstr" options but not in the goal
 
blanchet 
parents: 
41868 
diff
changeset
 | 
288  | 
val evals = evals |> map merge_tfrees  | 
| 41876 | 289  | 
val needs = needs |> Option.map (map merge_tfrees)  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
290  | 
val conj_ts = neg_t :: def_assm_ts @ nondef_assm_ts @ evals @ these needs  | 
| 33192 | 291  | 
val original_max_potential = max_potential  | 
292  | 
val original_max_genuine = max_genuine  | 
|
293  | 
val max_bisim_depth = fold Integer.max bisim_depths ~1  | 
|
| 55888 | 294  | 
val case_names = case_const_names ctxt  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
295  | 
val defs = def_assm_ts @ all_defs_of thy subst  | 
| 
42415
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
42361 
diff
changeset
 | 
296  | 
val nondefs = all_nondefs_of ctxt subst  | 
| 
41791
 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 
blanchet 
parents: 
41789 
diff
changeset
 | 
297  | 
val def_tables = const_def_tables ctxt subst defs  | 
| 
42415
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
42361 
diff
changeset
 | 
298  | 
val nondef_table = const_nondef_table nondefs  | 
| 35335 | 299  | 
val simp_table = Unsynchronized.ref (const_simp_table ctxt subst)  | 
300  | 
val psimp_table = const_psimp_table ctxt subst  | 
|
| 
35807
 
e4d1b5cbd429
added support for "specification" and "ax_specification" constructs to Nitpick
 
blanchet 
parents: 
35711 
diff
changeset
 | 
301  | 
val choice_spec_table = const_choice_spec_table ctxt subst  | 
| 
42415
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
42361 
diff
changeset
 | 
302  | 
val nondefs =  | 
| 
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
42361 
diff
changeset
 | 
303  | 
nondefs |> filter_out (is_choice_spec_axiom thy choice_spec_table)  | 
| 
41791
 
01d722707a36
always unfold constant defitions marked with "nitpick_def" -- to prevent unfolding, there's already "nitpick_simp"
 
blanchet 
parents: 
41789 
diff
changeset
 | 
304  | 
val intro_table = inductive_intro_table ctxt subst def_tables  | 
| 33192 | 305  | 
val ground_thm_table = ground_theorem_table thy  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
306  | 
val ersatz_table = ersatz_table ctxt  | 
| 
41007
 
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
 
blanchet 
parents: 
41001 
diff
changeset
 | 
307  | 
    val hol_ctxt as {wf_cache, ...} =
 | 
| 33192 | 308  | 
      {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes,
 | 
| 55888 | 309  | 
wfs = wfs, user_axioms = user_axioms, debug = debug, whacks = whacks,  | 
310  | 
binary_ints = binary_ints, destroy_constrs = destroy_constrs,  | 
|
311  | 
specialize = specialize, star_linear_preds = star_linear_preds,  | 
|
312  | 
total_consts = total_consts, needs = needs, tac_timeout = tac_timeout,  | 
|
313  | 
evals = evals, case_names = case_names, def_tables = def_tables,  | 
|
| 
42415
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
42361 
diff
changeset
 | 
314  | 
nondef_table = nondef_table, nondefs = nondefs, simp_table = simp_table,  | 
| 
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
42361 
diff
changeset
 | 
315  | 
psimp_table = psimp_table, choice_spec_table = choice_spec_table,  | 
| 
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
42361 
diff
changeset
 | 
316  | 
intro_table = intro_table, ground_thm_table = ground_thm_table,  | 
| 
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
42361 
diff
changeset
 | 
317  | 
ersatz_table = ersatz_table, skolems = Unsynchronized.ref [],  | 
| 
 
10accf397ab6
use "Spec_Rules" for finding axioms -- more reliable and cleaner
 
blanchet 
parents: 
42361 
diff
changeset
 | 
318  | 
special_funs = Unsynchronized.ref [],  | 
| 
33580
 
45c33e97cb86
added datatype constructor cache in Nitpick (to speed up the scope enumeration) and never test more than 4096 scopes
 
blanchet 
parents: 
33568 
diff
changeset
 | 
319  | 
unrolled_preds = Unsynchronized.ref [], wf_cache = Unsynchronized.ref [],  | 
| 
 
45c33e97cb86
added datatype constructor cache in Nitpick (to speed up the scope enumeration) and never test more than 4096 scopes
 
blanchet 
parents: 
33568 
diff
changeset
 | 
320  | 
constr_cache = Unsynchronized.ref []}  | 
| 
41789
 
7c7b68b06c1a
don't distinguish between "fixes" and other free variables -- this confuses users
 
blanchet 
parents: 
41772 
diff
changeset
 | 
321  | 
val pseudo_frees = []  | 
| 
41869
 
9e367f1c9570
robust handling of types occurring in "eval" and "preconstr" options but not in the goal
 
blanchet 
parents: 
41868 
diff
changeset
 | 
322  | 
val real_frees = fold Term.add_frees conj_ts []  | 
| 
 
9e367f1c9570
robust handling of types occurring in "eval" and "preconstr" options but not in the goal
 
blanchet 
parents: 
41868 
diff
changeset
 | 
323  | 
val _ = null (fold Term.add_tvars conj_ts []) orelse  | 
| 
 
9e367f1c9570
robust handling of types occurring in "eval" and "preconstr" options but not in the goal
 
blanchet 
parents: 
41868 
diff
changeset
 | 
324  | 
error "Nitpick cannot handle goals with schematic type variables."  | 
| 41875 | 325  | 
val (nondef_ts, def_ts, need_ts, got_all_mono_user_axioms,  | 
| 
41803
 
ef13e3b7cbaf
more work on "fix_datatype_vals" optimization (renamed "preconstruct")
 
blanchet 
parents: 
41802 
diff
changeset
 | 
326  | 
no_poly_user_axioms, binarize) =  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
327  | 
preprocess_formulas hol_ctxt nondef_assm_ts neg_t  | 
| 33192 | 328  | 
val got_all_user_axioms =  | 
329  | 
got_all_mono_user_axioms andalso no_poly_user_axioms  | 
|
330  | 
||
331  | 
fun print_wf (x, (gfp, wf)) =  | 
|
| 47560 | 332  | 
pprint_nt (fn () => Pretty.blk (0,  | 
| 33192 | 333  | 
          pstrs ("The " ^ (if gfp then "co" else "") ^ "inductive predicate \"")
 | 
334  | 
@ Syntax.pretty_term ctxt (Const x) ::  | 
|
335  | 
pstrs (if wf then  | 
|
336  | 
"\" was proved well-founded. Nitpick can compute it \  | 
|
337  | 
\efficiently."  | 
|
338  | 
else  | 
|
339  | 
"\" could not be proved well-founded. Nitpick might need to \  | 
|
340  | 
\unroll it.")))  | 
|
341  | 
val _ = if verbose then List.app print_wf (!wf_cache) else ()  | 
|
| 
38171
 
5f2ea92319e0
fix soundness bug w.r.t. "Suc" with "binary_ints"
 
blanchet 
parents: 
38170 
diff
changeset
 | 
342  | 
val das_wort_formula = if falsify then "Negated conjecture" else "Formula"  | 
| 33192 | 343  | 
val _ =  | 
| 
35386
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
344  | 
pprint_d (fn () => Pretty.chunks  | 
| 
38171
 
5f2ea92319e0
fix soundness bug w.r.t. "Suc" with "binary_ints"
 
blanchet 
parents: 
38170 
diff
changeset
 | 
345  | 
(pretties_for_formulas ctxt das_wort_formula [hd nondef_ts] @  | 
| 
35386
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
346  | 
pretties_for_formulas ctxt "Relevant definitional axiom" def_ts @  | 
| 
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
347  | 
pretties_for_formulas ctxt "Relevant nondefinitional axiom"  | 
| 
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
348  | 
(tl nondef_ts)))  | 
| 
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
349  | 
val _ = List.app (ignore o Term.type_of) (nondef_ts @ def_ts)  | 
| 33192 | 350  | 
handle TYPE (_, Ts, ts) =>  | 
351  | 
                   raise TYPE ("Nitpick.pick_them_nits_in_term", Ts, ts)
 | 
|
352  | 
||
| 
41801
 
ed77524f3429
first steps in implementing "fix_datatype_vals" optimization
 
blanchet 
parents: 
41793 
diff
changeset
 | 
353  | 
val nondef_us = nondef_ts |> map (nut_from_term hol_ctxt Eq)  | 
| 
 
ed77524f3429
first steps in implementing "fix_datatype_vals" optimization
 
blanchet 
parents: 
41793 
diff
changeset
 | 
354  | 
val def_us = def_ts |> map (nut_from_term hol_ctxt DefEq)  | 
| 41875 | 355  | 
val need_us = need_ts |> map (nut_from_term hol_ctxt Eq)  | 
| 
33558
 
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
 
blanchet 
parents: 
33556 
diff
changeset
 | 
356  | 
val (free_names, const_names) =  | 
| 
41888
 
79f837c2e33b
don't forget to look for constants appearing only in "need" option -- otherwise we get exceptions in "the_name" later on
 
blanchet 
parents: 
41876 
diff
changeset
 | 
357  | 
fold add_free_and_const_names (nondef_us @ def_us @ need_us) ([], [])  | 
| 
33558
 
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
 
blanchet 
parents: 
33556 
diff
changeset
 | 
358  | 
val (sel_names, nonsel_names) =  | 
| 
 
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
 
blanchet 
parents: 
33556 
diff
changeset
 | 
359  | 
List.partition (is_sel o nickname_of) const_names  | 
| 
46114
 
e6d33b200f42
remove subtlety whose justification got lost in time -- the new code is possibly less precise but sounder
 
blanchet 
parents: 
46103 
diff
changeset
 | 
360  | 
val sound_finitizes = none_true finitizes  | 
| 
33558
 
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
 
blanchet 
parents: 
33556 
diff
changeset
 | 
361  | 
(*  | 
| 
41888
 
79f837c2e33b
don't forget to look for constants appearing only in "need" option -- otherwise we get exceptions in "the_name" later on
 
blanchet 
parents: 
41876 
diff
changeset
 | 
362  | 
val _ =  | 
| 
 
79f837c2e33b
don't forget to look for constants appearing only in "need" option -- otherwise we get exceptions in "the_name" later on
 
blanchet 
parents: 
41876 
diff
changeset
 | 
363  | 
List.app (print_g o string_for_nut ctxt) (nondef_us @ def_us @ need_us)  | 
| 
33558
 
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
 
blanchet 
parents: 
33556 
diff
changeset
 | 
364  | 
*)  | 
| 
 
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
 
blanchet 
parents: 
33556 
diff
changeset
 | 
365  | 
|
| 
34121
 
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
 
blanchet 
parents: 
34039 
diff
changeset
 | 
366  | 
val unique_scope = forall (curry (op =) 1 o length o snd) cards_assigns  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
367  | 
fun monotonicity_message Ts extra =  | 
| 38188 | 368  | 
let val pretties = map (pretty_maybe_quote o pretty_for_type ctxt) Ts in  | 
369  | 
Pretty.blk (0,  | 
|
370  | 
          pstrs ("The type" ^ plural_s_for_list pretties ^ " ") @
 | 
|
371  | 
pretty_serial_commas "and" pretties @  | 
|
372  | 
          pstrs (" " ^
 | 
|
| 42959 | 373  | 
(if none_true monos then  | 
374  | 
"passed the monotonicity test"  | 
|
375  | 
else  | 
|
376  | 
(if length pretties = 1 then "is" else "are") ^  | 
|
377  | 
" considered monotonic") ^  | 
|
378  | 
". " ^ extra))  | 
|
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
379  | 
end  | 
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
380  | 
fun is_type_fundamentally_monotonic T =  | 
| 55890 | 381  | 
(is_data_type ctxt T andalso not (is_quot_type ctxt T) andalso  | 
| 
37256
 
0dca1ec52999
thread along context instead of theory for typedef lookup
 
blanchet 
parents: 
37213 
diff
changeset
 | 
382  | 
(not (is_pure_typedef ctxt T) orelse is_univ_typedef ctxt T)) orelse  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38214 
diff
changeset
 | 
383  | 
is_number_type ctxt T orelse is_bit_type T  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
384  | 
fun is_type_actually_monotonic T =  | 
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
385  | 
TimeLimit.timeLimit tac_timeout (formulas_monotonic hol_ctxt binarize T)  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
386  | 
(nondef_ts, def_ts)  | 
| 47716 | 387  | 
handle TimeLimit.TimeOut => false  | 
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
388  | 
fun is_type_kind_of_monotonic T =  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
389  | 
case triple_lookup (type_match thy) monos T of  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
390  | 
SOME (SOME false) => false  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
391  | 
| _ => is_type_actually_monotonic T  | 
| 
34936
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
392  | 
fun is_type_monotonic T =  | 
| 33192 | 393  | 
unique_scope orelse  | 
394  | 
case triple_lookup (type_match thy) monos T of  | 
|
395  | 
SOME (SOME b) => b  | 
|
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
396  | 
| _ => is_type_fundamentally_monotonic T orelse  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
397  | 
is_type_actually_monotonic T  | 
| 55890 | 398  | 
fun is_deep_data_type_finitizable T =  | 
| 
41793
 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 
blanchet 
parents: 
41791 
diff
changeset
 | 
399  | 
triple_lookup (type_match thy) finitizes T = SOME (SOME true)  | 
| 55890 | 400  | 
    fun is_shallow_data_type_finitizable (T as Type (@{type_name fun_box}, _)) =
 | 
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
401  | 
is_type_kind_of_monotonic T  | 
| 55890 | 402  | 
| is_shallow_data_type_finitizable T =  | 
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
403  | 
case triple_lookup (type_match thy) finitizes T of  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
404  | 
SOME (SOME b) => b  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
405  | 
| _ => is_type_kind_of_monotonic T  | 
| 55890 | 406  | 
fun is_data_type_deep T =  | 
| 55888 | 407  | 
      T = @{typ unit} orelse T = nat_T orelse is_word_type T orelse
 | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
408  | 
exists (curry (op =) T o domain_type o type_of) sel_names  | 
| 
35386
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
409  | 
val all_Ts = ground_types_in_terms hol_ctxt binarize (nondef_ts @ def_ts)  | 
| 35408 | 410  | 
|> sort Term_Ord.typ_ord  | 
| 38214 | 411  | 
val _ =  | 
412  | 
if verbose andalso binary_ints = SOME true andalso  | 
|
413  | 
exists (member (op =) [nat_T, int_T]) all_Ts then  | 
|
414  | 
print_v (K "The option \"binary_ints\" will be ignored because of the \  | 
|
415  | 
\presence of rationals, reals, \"Suc\", \"gcd\", or \"lcm\" \  | 
|
| 55888 | 416  | 
\in the problem.")  | 
| 38214 | 417  | 
else  | 
418  | 
()  | 
|
419  | 
val _ =  | 
|
| 43022 | 420  | 
if mode = Normal andalso  | 
| 38214 | 421  | 
         exists (fn Type (@{type_name Datatype.node}, _) => true | _ => false)
 | 
422  | 
all_Ts then  | 
|
| 47560 | 423  | 
        print_nt (K ("Warning: The problem involves directly or indirectly the \
 | 
424  | 
                     \internal type " ^ quote @{type_name Datatype.node} ^
 | 
|
425  | 
". This type is very Nitpick-unfriendly, and its presence \  | 
|
426  | 
\usually indicates either a failure of abstraction or a \  | 
|
427  | 
\quirk in Nitpick."))  | 
|
| 38214 | 428  | 
else  | 
429  | 
()  | 
|
| 
34982
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
430  | 
val (mono_Ts, nonmono_Ts) = List.partition is_type_monotonic all_Ts  | 
| 33192 | 431  | 
val _ =  | 
| 
34936
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
432  | 
if verbose andalso not unique_scope then  | 
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35408 
diff
changeset
 | 
433  | 
case filter_out is_type_fundamentally_monotonic mono_Ts of  | 
| 
34936
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
434  | 
[] => ()  | 
| 
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
435  | 
| interesting_mono_Ts =>  | 
| 38188 | 436  | 
pprint_v (K (monotonicity_message interesting_mono_Ts  | 
437  | 
"Nitpick might be able to skip some scopes."))  | 
|
| 33192 | 438  | 
else  | 
439  | 
()  | 
|
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
440  | 
val (deep_dataTs, shallow_dataTs) =  | 
| 55890 | 441  | 
all_Ts |> filter (is_data_type ctxt)  | 
442  | 
|> List.partition is_data_type_deep  | 
|
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
443  | 
val finitizable_dataTs =  | 
| 
41793
 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 
blanchet 
parents: 
41791 
diff
changeset
 | 
444  | 
(deep_dataTs |> filter_out (is_finite_type hol_ctxt)  | 
| 55890 | 445  | 
|> filter is_deep_data_type_finitizable) @  | 
| 
41793
 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 
blanchet 
parents: 
41791 
diff
changeset
 | 
446  | 
(shallow_dataTs |> filter_out (is_finite_type hol_ctxt)  | 
| 55890 | 447  | 
|> filter is_shallow_data_type_finitizable)  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
448  | 
val _ =  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
449  | 
if verbose andalso not (null finitizable_dataTs) then  | 
| 38188 | 450  | 
pprint_v (K (monotonicity_message finitizable_dataTs  | 
451  | 
"Nitpick can use a more precise finite encoding."))  | 
|
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
452  | 
else  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35384 
diff
changeset
 | 
453  | 
()  | 
| 33192 | 454  | 
(*  | 
| 
36128
 
a3d8d5329438
make Nitpick output everything to tracing in debug mode;
 
blanchet 
parents: 
36126 
diff
changeset
 | 
455  | 
val _ = print_g "Monotonic types:"  | 
| 
 
a3d8d5329438
make Nitpick output everything to tracing in debug mode;
 
blanchet 
parents: 
36126 
diff
changeset
 | 
456  | 
val _ = List.app (print_g o string_for_type ctxt) mono_Ts  | 
| 
 
a3d8d5329438
make Nitpick output everything to tracing in debug mode;
 
blanchet 
parents: 
36126 
diff
changeset
 | 
457  | 
val _ = print_g "Nonmonotonic types:"  | 
| 
 
a3d8d5329438
make Nitpick output everything to tracing in debug mode;
 
blanchet 
parents: 
36126 
diff
changeset
 | 
458  | 
val _ = List.app (print_g o string_for_type ctxt) nonmono_Ts  | 
| 33192 | 459  | 
*)  | 
460  | 
||
| 36384 | 461  | 
val incremental = Int.max (max_potential, max_genuine) >= 2  | 
462  | 
val actual_sat_solver =  | 
|
| 33192 | 463  | 
if sat_solver <> "smart" then  | 
| 36384 | 464  | 
if incremental andalso  | 
| 
34936
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
465  | 
not (member (op =) (Kodkod_SAT.configured_sat_solvers true)  | 
| 
35333
 
f61de25f71f9
distinguish between Kodkodi warnings and errors in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
466  | 
sat_solver) then  | 
| 47560 | 467  | 
          (print_nt (K ("An incremental SAT solver is required: \"SAT4J\" will \
 | 
468  | 
\be used instead of " ^ quote sat_solver ^ "."));  | 
|
| 33192 | 469  | 
"SAT4J")  | 
470  | 
else  | 
|
471  | 
sat_solver  | 
|
472  | 
else  | 
|
| 36384 | 473  | 
Kodkod_SAT.smart_sat_solver_name incremental  | 
| 33192 | 474  | 
val _ =  | 
475  | 
if sat_solver = "smart" then  | 
|
| 36384 | 476  | 
print_v (fn () =>  | 
477  | 
"Using SAT solver " ^ quote actual_sat_solver ^ ". The following" ^  | 
|
478  | 
(if incremental then " incremental " else " ") ^  | 
|
479  | 
"solvers are configured: " ^  | 
|
480  | 
commas_quote (Kodkod_SAT.configured_sat_solvers incremental) ^ ".")  | 
|
| 33192 | 481  | 
else  | 
482  | 
()  | 
|
483  | 
||
484  | 
val too_big_scopes = Unsynchronized.ref []  | 
|
485  | 
||
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
486  | 
fun problem_for_scope unsound  | 
| 55890 | 487  | 
            (scope as {card_assigns, bits, bisim_depth, data_types, ofs, ...}) =
 | 
| 33192 | 488  | 
let  | 
489  | 
val _ = not (exists (fn other => scope_less_eq other scope)  | 
|
| 
34936
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
490  | 
(!too_big_scopes)) orelse  | 
| 
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
491  | 
                raise TOO_LARGE ("Nitpick.pick_them_nits_in_term.\
 | 
| 
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
492  | 
\problem_for_scope", "too large scope")  | 
| 33192 | 493  | 
(*  | 
| 
36128
 
a3d8d5329438
make Nitpick output everything to tracing in debug mode;
 
blanchet 
parents: 
36126 
diff
changeset
 | 
494  | 
val _ = print_g "Offsets:"  | 
| 33192 | 495  | 
val _ = List.app (fn (T, j0) =>  | 
| 
36128
 
a3d8d5329438
make Nitpick output everything to tracing in debug mode;
 
blanchet 
parents: 
36126 
diff
changeset
 | 
496  | 
print_g (string_for_type ctxt T ^ " = " ^  | 
| 
 
a3d8d5329438
make Nitpick output everything to tracing in debug mode;
 
blanchet 
parents: 
36126 
diff
changeset
 | 
497  | 
string_of_int j0))  | 
| 33192 | 498  | 
(Typtab.dest ofs)  | 
499  | 
*)  | 
|
| 
41858
 
37ce158d6266
if "total_consts" is set, report cex's as quasi-genuine
 
blanchet 
parents: 
41856 
diff
changeset
 | 
500  | 
val effective_total_consts =  | 
| 41856 | 501  | 
case total_consts of  | 
502  | 
SOME b => b  | 
|
| 55890 | 503  | 
| NONE => forall (is_exact_type data_types true) all_Ts  | 
| 33192 | 504  | 
val main_j0 = offset_of_type ofs bool_T  | 
505  | 
val (nat_card, nat_j0) = spec_of_type scope nat_T  | 
|
506  | 
val (int_card, int_j0) = spec_of_type scope int_T  | 
|
| 
34936
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
507  | 
val _ = (nat_j0 = main_j0 andalso int_j0 = main_j0) orelse  | 
| 
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
508  | 
                raise BAD ("Nitpick.pick_them_nits_in_term.problem_for_scope",
 | 
| 
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
509  | 
"bad offsets")  | 
| 33192 | 510  | 
val kk = kodkod_constrs peephole_optim nat_card int_card main_j0  | 
511  | 
val (free_names, rep_table) =  | 
|
| 38170 | 512  | 
choose_reps_for_free_vars scope pseudo_frees free_names  | 
513  | 
NameTable.empty  | 
|
| 33192 | 514  | 
val (sel_names, rep_table) = choose_reps_for_all_sels scope rep_table  | 
| 38170 | 515  | 
val (nonsel_names, rep_table) =  | 
| 
41858
 
37ce158d6266
if "total_consts" is set, report cex's as quasi-genuine
 
blanchet 
parents: 
41856 
diff
changeset
 | 
516  | 
choose_reps_for_consts scope effective_total_consts nonsel_names  | 
| 
 
37ce158d6266
if "total_consts" is set, report cex's as quasi-genuine
 
blanchet 
parents: 
41856 
diff
changeset
 | 
517  | 
rep_table  | 
| 38182 | 518  | 
val (guiltiest_party, min_highest_arity) =  | 
519  | 
NameTable.fold (fn (name, R) => fn (s, n) =>  | 
|
520  | 
let val n' = arity_of_rep R in  | 
|
521  | 
if n' > n then (nickname_of name, n') else (s, n)  | 
|
522  | 
                             end) rep_table ("", 1)
 | 
|
| 33192 | 523  | 
val min_univ_card =  | 
| 36384 | 524  | 
NameTable.fold (Integer.max o min_univ_card_of_rep o snd) rep_table  | 
| 34126 | 525  | 
(univ_card nat_card int_card main_j0 [] KK.True)  | 
| 52202 | 526  | 
val _ = check_arity guiltiest_party min_univ_card min_highest_arity  | 
| 33192 | 527  | 
|
| 41802 | 528  | 
val def_us =  | 
529  | 
def_us |> map (choose_reps_in_nut scope unsound rep_table true)  | 
|
530  | 
val nondef_us =  | 
|
531  | 
nondef_us |> map (choose_reps_in_nut scope unsound rep_table false)  | 
|
| 41875 | 532  | 
val need_us =  | 
533  | 
need_us |> map (choose_reps_in_nut scope unsound rep_table false)  | 
|
| 33745 | 534  | 
(*  | 
| 
36128
 
a3d8d5329438
make Nitpick output everything to tracing in debug mode;
 
blanchet 
parents: 
36126 
diff
changeset
 | 
535  | 
val _ = List.app (print_g o string_for_nut ctxt)  | 
| 33192 | 536  | 
(free_names @ sel_names @ nonsel_names @  | 
| 
41888
 
79f837c2e33b
don't forget to look for constants appearing only in "need" option -- otherwise we get exceptions in "the_name" later on
 
blanchet 
parents: 
41876 
diff
changeset
 | 
537  | 
nondef_us @ def_us @ need_us)  | 
| 33745 | 538  | 
*)  | 
| 33192 | 539  | 
val (free_rels, pool, rel_table) =  | 
540  | 
rename_free_vars free_names initial_pool NameTable.empty  | 
|
541  | 
val (sel_rels, pool, rel_table) =  | 
|
542  | 
rename_free_vars sel_names pool rel_table  | 
|
543  | 
val (other_rels, pool, rel_table) =  | 
|
544  | 
rename_free_vars nonsel_names pool rel_table  | 
|
| 41802 | 545  | 
val nondef_us = nondef_us |> map (rename_vars_in_nut pool rel_table)  | 
546  | 
val def_us = def_us |> map (rename_vars_in_nut pool rel_table)  | 
|
| 41875 | 547  | 
val need_us = need_us |> map (rename_vars_in_nut pool rel_table)  | 
| 
35386
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
548  | 
val nondef_fs = map (kodkod_formula_from_nut ofs kk) nondef_us  | 
| 
35280
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
35220 
diff
changeset
 | 
549  | 
val def_fs = map (kodkod_formula_from_nut ofs kk) def_us  | 
| 
35386
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
550  | 
val formula = fold (fold s_and) [def_fs, nondef_fs] KK.True  | 
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
551  | 
val comment = (if unsound then "unsound" else "sound") ^ "\n" ^  | 
| 
37146
 
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
 
wenzelm 
parents: 
36913 
diff
changeset
 | 
552  | 
Print_Mode.setmp [] multiline_string_for_scope scope  | 
| 34998 | 553  | 
val kodkod_sat_solver =  | 
| 36384 | 554  | 
Kodkod_SAT.sat_solver_spec actual_sat_solver |> snd  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
555  | 
val bit_width = if bits = 0 then 16 else bits + 1  | 
| 36384 | 556  | 
val delay =  | 
557  | 
if unsound then  | 
|
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
558  | 
unsound_delay_for_timeout (Time.- (deadline, Time.now ()))  | 
| 36384 | 559  | 
else  | 
560  | 
0  | 
|
561  | 
        val settings = [("solver", commas_quote kodkod_sat_solver),
 | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
562  | 
                        ("bit_width", string_of_int bit_width),
 | 
| 38124 | 563  | 
                        ("symmetry_breaking", string_of_int kodkod_sym_break),
 | 
| 
36386
 
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
 
blanchet 
parents: 
36385 
diff
changeset
 | 
564  | 
                        ("sharing", "3"),
 | 
| 
 
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
 
blanchet 
parents: 
36385 
diff
changeset
 | 
565  | 
                        ("flatten", "false"),
 | 
| 33192 | 566  | 
                        ("delay", signed_string_of_int delay)]
 | 
567  | 
val plain_rels = free_rels @ other_rels  | 
|
568  | 
val plain_bounds = map (bound_for_plain_rel ctxt debug) plain_rels  | 
|
569  | 
val plain_axioms = map (declarative_axiom_for_plain_rel kk) plain_rels  | 
|
| 41995 | 570  | 
val need_vals =  | 
571  | 
          map (fn dtype as {typ, ...} =>
 | 
|
| 55890 | 572  | 
(typ, needed_values_for_data_type need_us ofs dtype))  | 
573  | 
data_types  | 
|
| 41995 | 574  | 
val sel_bounds =  | 
| 55890 | 575  | 
map (bound_for_sel_rel ctxt debug need_vals data_types) sel_rels  | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35185 
diff
changeset
 | 
576  | 
val dtype_axioms =  | 
| 55890 | 577  | 
declarative_axioms_for_data_types hol_ctxt binarize need_us need_vals  | 
578  | 
datatype_sym_break bits ofs kk rel_table data_types  | 
|
| 33192 | 579  | 
val declarative_axioms = plain_axioms @ dtype_axioms  | 
| 
35220
 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 
blanchet 
parents: 
35190 
diff
changeset
 | 
580  | 
val univ_card = Int.max (univ_card nat_card int_card main_j0  | 
| 
 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 
blanchet 
parents: 
35190 
diff
changeset
 | 
581  | 
(plain_bounds @ sel_bounds) formula,  | 
| 
 
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
 
blanchet 
parents: 
35190 
diff
changeset
 | 
582  | 
main_j0 |> bits > 0 ? Integer.add (bits + 1))  | 
| 38126 | 583  | 
val (built_in_bounds, built_in_axioms) =  | 
| 39345 | 584  | 
bounds_and_axioms_for_built_in_rels_in_formulas debug univ_card  | 
585  | 
nat_card int_card main_j0 (formula :: declarative_axioms)  | 
|
| 33192 | 586  | 
val bounds = built_in_bounds @ plain_bounds @ sel_bounds  | 
587  | 
|> not debug ? merge_bounds  | 
|
| 38126 | 588  | 
val axioms = built_in_axioms @ declarative_axioms  | 
| 33192 | 589  | 
val highest_arity =  | 
590  | 
fold Integer.max (map (fst o fst) (maps fst bounds)) 0  | 
|
| 38126 | 591  | 
val formula = fold_rev s_and axioms formula  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
592  | 
val _ = if bits = 0 then () else check_bits bits formula  | 
| 52202 | 593  | 
val _ = if formula = KK.False then ()  | 
| 38182 | 594  | 
else check_arity "" univ_card highest_arity  | 
| 33192 | 595  | 
in  | 
596  | 
        SOME ({comment = comment, settings = settings, univ_card = univ_card,
 | 
|
597  | 
tuple_assigns = [], bounds = bounds,  | 
|
| 
35280
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
35220 
diff
changeset
 | 
598  | 
int_bounds = if bits = 0 then sequential_int_bounds univ_card  | 
| 
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
35220 
diff
changeset
 | 
599  | 
else pow_of_two_int_bounds bits main_j0,  | 
| 33192 | 600  | 
expr_assigns = [], formula = formula},  | 
601  | 
              {free_names = free_names, sel_names = sel_names,
 | 
|
602  | 
nonsel_names = nonsel_names, rel_table = rel_table,  | 
|
| 
35386
 
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
 
blanchet 
parents: 
35385 
diff
changeset
 | 
603  | 
unsound = unsound, scope = scope})  | 
| 33192 | 604  | 
end  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
605  | 
handle TOO_LARGE (loc, msg) =>  | 
| 
34936
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
606  | 
if loc = "Nitpick_Kodkod.check_arity" andalso  | 
| 
 
c4f04bee79f3
some work on Nitpick's support for quotient types;
 
blanchet 
parents: 
34935 
diff
changeset
 | 
607  | 
not (Typtab.is_empty ofs) then  | 
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
608  | 
problem_for_scope unsound  | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35185 
diff
changeset
 | 
609  | 
                   {hol_ctxt = hol_ctxt, binarize = binarize,
 | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35185 
diff
changeset
 | 
610  | 
card_assigns = card_assigns, bits = bits,  | 
| 55890 | 611  | 
bisim_depth = bisim_depth, data_types = data_types,  | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35185 
diff
changeset
 | 
612  | 
ofs = Typtab.empty}  | 
| 33192 | 613  | 
else if loc = "Nitpick.pick_them_nits_in_term.\  | 
614  | 
\problem_for_scope" then  | 
|
615  | 
NONE  | 
|
616  | 
else  | 
|
617  | 
(Unsynchronized.change too_big_scopes (cons scope);  | 
|
| 
41992
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
618  | 
print_v (fn () =>  | 
| 
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
619  | 
"Limit reached: " ^ msg ^ ". Skipping " ^  | 
| 
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
620  | 
(if unsound then "potential component of " else "") ^  | 
| 
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
621  | 
"scope.");  | 
| 33192 | 622  | 
NONE)  | 
| 
35280
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
35220 
diff
changeset
 | 
623  | 
| TOO_SMALL (_, msg) =>  | 
| 
41992
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
624  | 
(print_v (fn () =>  | 
| 
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
625  | 
"Limit reached: " ^ msg ^ ". Skipping " ^  | 
| 
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
626  | 
(if unsound then "potential component of " else "") ^  | 
| 
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
627  | 
"scope.");  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
628  | 
NONE)  | 
| 33192 | 629  | 
|
| 55888 | 630  | 
val das_wort_model = if falsify then "counterexample" else "model"  | 
| 33192 | 631  | 
|
632  | 
val scopes = Unsynchronized.ref []  | 
|
633  | 
val generated_scopes = Unsynchronized.ref []  | 
|
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
634  | 
val generated_problems = Unsynchronized.ref ([] : rich_problem list)  | 
| 33192 | 635  | 
val checked_problems = Unsynchronized.ref (SOME [])  | 
636  | 
val met_potential = Unsynchronized.ref 0  | 
|
637  | 
||
638  | 
fun update_checked_problems problems =  | 
|
639  | 
List.app (Unsynchronized.change checked_problems o Option.map o cons  | 
|
640  | 
o nth problems)  | 
|
| 
35333
 
f61de25f71f9
distinguish between Kodkodi warnings and errors in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
641  | 
fun show_kodkod_warning "" = ()  | 
| 47560 | 642  | 
| show_kodkod_warning s = print_nt (fn () => "Kodkod warning: " ^ s ^ ".")  | 
| 33192 | 643  | 
|
644  | 
fun print_and_check_model genuine bounds  | 
|
645  | 
            ({free_names, sel_names, nonsel_names, rel_table, scope, ...}
 | 
|
646  | 
: problem_extension) =  | 
|
647  | 
let  | 
|
| 47752 | 648  | 
val _ =  | 
649  | 
print_t (fn () =>  | 
|
650  | 
"% SZS status " ^  | 
|
651  | 
(if genuine then  | 
|
652  | 
if falsify then "CounterSatisfiable" else "Satisfiable"  | 
|
653  | 
else  | 
|
654  | 
"Unknown") ^ "\n" ^  | 
|
655  | 
"% SZS output start FiniteModel")  | 
|
| 33192 | 656  | 
val (reconstructed_model, codatatypes_ok) =  | 
| 
41993
 
bd6296de1432
reintroduced "show_skolems" option -- useful when too many Skolems are displayed
 
blanchet 
parents: 
41992 
diff
changeset
 | 
657  | 
reconstruct_hol_model  | 
| 55889 | 658  | 
              {show_types = show_types, show_skolems = show_skolems,
 | 
| 
41993
 
bd6296de1432
reintroduced "show_skolems" option -- useful when too many Skolems are displayed
 
blanchet 
parents: 
41992 
diff
changeset
 | 
659  | 
show_consts = show_consts}  | 
| 38170 | 660  | 
scope formats atomss real_frees pseudo_frees free_names sel_names  | 
661  | 
nonsel_names rel_table bounds  | 
|
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
662  | 
val genuine_means_genuine =  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
663  | 
got_all_user_axioms andalso none_true wfs andalso  | 
| 
41858
 
37ce158d6266
if "total_consts" is set, report cex's as quasi-genuine
 
blanchet 
parents: 
41856 
diff
changeset
 | 
664  | 
sound_finitizes andalso total_consts <> SOME true andalso  | 
| 
 
37ce158d6266
if "total_consts" is set, report cex's as quasi-genuine
 
blanchet 
parents: 
41856 
diff
changeset
 | 
665  | 
codatatypes_ok  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
666  | 
fun assms_prop () =  | 
| 
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
667  | 
Logic.mk_conjunction_list (neg_t :: def_assm_ts @ nondef_assm_ts)  | 
| 33192 | 668  | 
in  | 
| 47752 | 669  | 
(pprint_a (Pretty.chunks  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
670  | 
[Pretty.blk (0,  | 
| 43022 | 671  | 
(pstrs ((if mode = Auto_Try then "Auto " else "") ^  | 
672  | 
"Nitpick found a" ^  | 
|
| 
41992
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
673  | 
(if not genuine then " potentially spurious "  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
674  | 
else if genuine_means_genuine then " "  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
675  | 
else " quasi genuine ") ^ das_wort_model) @  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
676  | 
(case pretties_for_scope scope verbose of  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
677  | 
[] => []  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
678  | 
| pretties => pstrs " for " @ pretties) @  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
679  | 
[Pretty.str ":\n"])),  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
680  | 
Pretty.indent indent_size reconstructed_model]);  | 
| 47564 | 681  | 
print_t (K "% SZS output end FiniteModel");  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
682  | 
if genuine then  | 
| 55888 | 683  | 
(if check_genuine then  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
684  | 
case prove_hol_model scope tac_timeout free_names sel_names  | 
| 
38169
 
b51784515471
optimize local "def"s by treating them as definitions
 
blanchet 
parents: 
38126 
diff
changeset
 | 
685  | 
rel_table bounds (assms_prop ()) of  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
686  | 
SOME true =>  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
687  | 
                print ("Confirmation by \"auto\": The above " ^
 | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
688  | 
das_wort_model ^ " is really genuine.")  | 
| 33192 | 689  | 
| SOME false =>  | 
| 
34982
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
690  | 
if genuine_means_genuine then  | 
| 
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
691  | 
                  error ("A supposedly genuine " ^ das_wort_model ^ " was \
 | 
| 
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
692  | 
\shown to be spurious by \"auto\".\nThis should never \  | 
| 
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
693  | 
\happen.\nPlease send a bug report to blanchet\  | 
| 33192 | 694  | 
\te@in.tum.de.")  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
695  | 
else  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
696  | 
                   print ("Refutation by \"auto\": The above " ^
 | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
697  | 
das_wort_model ^ " is spurious.")  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
698  | 
| NONE => print "No confirmation by \"auto\"."  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
699  | 
else  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
700  | 
();  | 
| 
41048
 
d5ebe94248ad
added a hint when the user obviously just forgot a colon after the lemma's name
 
blanchet 
parents: 
41007 
diff
changeset
 | 
701  | 
if has_lonely_bool_var orig_t then  | 
| 
 
d5ebe94248ad
added a hint when the user obviously just forgot a colon after the lemma's name
 
blanchet 
parents: 
41007 
diff
changeset
 | 
702  | 
print "Hint: Maybe you forgot a colon after the lemma's name?"  | 
| 
 
d5ebe94248ad
added a hint when the user obviously just forgot a colon after the lemma's name
 
blanchet 
parents: 
41007 
diff
changeset
 | 
703  | 
else if has_syntactic_sorts orig_t then  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
704  | 
print "Hint: Maybe you forgot a type constraint?"  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
705  | 
else  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
706  | 
();  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
707  | 
if not genuine_means_genuine then  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
708  | 
if no_poly_user_axioms then  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
709  | 
let  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
710  | 
val options =  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
711  | 
[] |> not got_all_mono_user_axioms  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
712  | 
                          ? cons ("user_axioms", "\"true\"")
 | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
713  | 
|> not (none_true wfs)  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
714  | 
                          ? cons ("wf", "\"smart\" or \"false\"")
 | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
715  | 
|> not sound_finitizes  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
716  | 
                          ? cons ("finitize", "\"smart\" or \"false\"")
 | 
| 
41858
 
37ce158d6266
if "total_consts" is set, report cex's as quasi-genuine
 
blanchet 
parents: 
41856 
diff
changeset
 | 
717  | 
|> total_consts = SOME true  | 
| 
 
37ce158d6266
if "total_consts" is set, report cex's as quasi-genuine
 
blanchet 
parents: 
41856 
diff
changeset
 | 
718  | 
                          ? cons ("total_consts", "\"smart\" or \"false\"")
 | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
719  | 
|> not codatatypes_ok  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
720  | 
                          ? cons ("bisim_depth", "a nonnegative value")
 | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
721  | 
val ss =  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
722  | 
map (fn (name, value) => quote name ^ " set to " ^ value)  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
723  | 
options  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
724  | 
in  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
725  | 
                  print ("Try again with " ^
 | 
| 41872 | 726  | 
space_implode " " (serial_commas "and" ss) ^  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
727  | 
" to confirm that the " ^ das_wort_model ^  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
728  | 
" is genuine.")  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
729  | 
end  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
730  | 
else  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
731  | 
                print ("Nitpick is unable to guarantee the authenticity of \
 | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
732  | 
\the " ^ das_wort_model ^ " in the presence of \  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
733  | 
\polymorphic axioms.")  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
734  | 
else  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
735  | 
();  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
736  | 
NONE)  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
737  | 
else  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
738  | 
if not genuine then  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
739  | 
(Unsynchronized.inc met_potential;  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
740  | 
if check_potential then  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
741  | 
let  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
742  | 
val status = prove_hol_model scope tac_timeout free_names  | 
| 
38169
 
b51784515471
optimize local "def"s by treating them as definitions
 
blanchet 
parents: 
38126 
diff
changeset
 | 
743  | 
sel_names rel_table bounds  | 
| 
 
b51784515471
optimize local "def"s by treating them as definitions
 
blanchet 
parents: 
38126 
diff
changeset
 | 
744  | 
(assms_prop ())  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
745  | 
in  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
746  | 
(case status of  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
747  | 
                     SOME true => print ("Confirmation by \"auto\": The \
 | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
748  | 
\above " ^ das_wort_model ^  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
749  | 
" is genuine.")  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
750  | 
                   | SOME false => print ("Refutation by \"auto\": The above " ^
 | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
751  | 
das_wort_model ^ " is spurious.")  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
752  | 
| NONE => print "No confirmation by \"auto\".");  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
753  | 
status  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
754  | 
end  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
755  | 
else  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
756  | 
NONE)  | 
| 33192 | 757  | 
else  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
758  | 
NONE)  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
759  | 
|> pair genuine_means_genuine  | 
| 33192 | 760  | 
end  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
761  | 
fun solve_any_problem (found_really_genuine, max_potential, max_genuine,  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
762  | 
donno) first_time problems =  | 
| 33192 | 763  | 
let  | 
764  | 
val max_potential = Int.max (0, max_potential)  | 
|
765  | 
val max_genuine = Int.max (0, max_genuine)  | 
|
766  | 
fun print_and_check genuine (j, bounds) =  | 
|
767  | 
print_and_check_model genuine bounds (snd (nth problems j))  | 
|
768  | 
val max_solutions = max_potential + max_genuine  | 
|
| 36384 | 769  | 
|> not incremental ? Integer.min 1  | 
| 33192 | 770  | 
in  | 
771  | 
if max_solutions <= 0 then  | 
|
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
772  | 
(found_really_genuine, 0, 0, donno)  | 
| 33192 | 773  | 
else  | 
| 
41793
 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 
blanchet 
parents: 
41791 
diff
changeset
 | 
774  | 
case KK.solve_any_problem debug overlord deadline max_threads  | 
| 
 
c7a2669ae75d
tweaked Nitpick based on C++ memory model example
 
blanchet 
parents: 
41791 
diff
changeset
 | 
775  | 
max_solutions (map fst problems) of  | 
| 49024 | 776  | 
KK.JavaNotFound =>  | 
| 55889 | 777  | 
(print_nt (K java_not_found_message);  | 
| 
35696
 
17ae461d6133
show nice error message in Nitpick when "java" is not available
 
blanchet 
parents: 
35671 
diff
changeset
 | 
778  | 
(found_really_genuine, max_potential, max_genuine, donno + 1))  | 
| 
38516
 
307669429dc1
gracefully handle the case where the JVM is too old in Nitpick
 
blanchet 
parents: 
38240 
diff
changeset
 | 
779  | 
| KK.JavaTooOld =>  | 
| 55889 | 780  | 
(print_nt (K java_too_old_message);  | 
| 
38516
 
307669429dc1
gracefully handle the case where the JVM is too old in Nitpick
 
blanchet 
parents: 
38240 
diff
changeset
 | 
781  | 
(found_really_genuine, max_potential, max_genuine, donno + 1))  | 
| 
35696
 
17ae461d6133
show nice error message in Nitpick when "java" is not available
 
blanchet 
parents: 
35671 
diff
changeset
 | 
782  | 
| KK.KodkodiNotInstalled =>  | 
| 55889 | 783  | 
(print_nt (K kodkodi_not_installed_message);  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
784  | 
(found_really_genuine, max_potential, max_genuine, donno + 1))  | 
| 50487 | 785  | 
| KK.KodkodiTooOld =>  | 
| 55889 | 786  | 
(print_nt (K kodkodi_too_old_message);  | 
| 50487 | 787  | 
(found_really_genuine, max_potential, max_genuine, donno + 1))  | 
| 
35333
 
f61de25f71f9
distinguish between Kodkodi warnings and errors in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
788  | 
| KK.Normal ([], unsat_js, s) =>  | 
| 
 
f61de25f71f9
distinguish between Kodkodi warnings and errors in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
789  | 
(update_checked_problems problems unsat_js; show_kodkod_warning s;  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
790  | 
(found_really_genuine, max_potential, max_genuine, donno))  | 
| 
35333
 
f61de25f71f9
distinguish between Kodkodi warnings and errors in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
791  | 
| KK.Normal (sat_ps, unsat_js, s) =>  | 
| 33192 | 792  | 
let  | 
793  | 
val (lib_ps, con_ps) =  | 
|
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
794  | 
List.partition (#unsound o snd o nth problems o fst) sat_ps  | 
| 33192 | 795  | 
in  | 
796  | 
update_checked_problems problems (unsat_js @ map fst lib_ps);  | 
|
| 
35333
 
f61de25f71f9
distinguish between Kodkodi warnings and errors in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
797  | 
show_kodkod_warning s;  | 
| 33192 | 798  | 
if null con_ps then  | 
799  | 
let  | 
|
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
800  | 
val genuine_codes =  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
801  | 
lib_ps |> take max_potential  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
802  | 
|> map (print_and_check false)  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
803  | 
|> filter (curry (op =) (SOME true) o snd)  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
804  | 
val found_really_genuine =  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
805  | 
found_really_genuine orelse exists fst genuine_codes  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
806  | 
val num_genuine = length genuine_codes  | 
| 33192 | 807  | 
val max_genuine = max_genuine - num_genuine  | 
808  | 
val max_potential = max_potential  | 
|
809  | 
- (length lib_ps - num_genuine)  | 
|
810  | 
in  | 
|
811  | 
if max_genuine <= 0 then  | 
|
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
812  | 
(found_really_genuine, 0, 0, donno)  | 
| 33192 | 813  | 
else  | 
814  | 
let  | 
|
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
815  | 
(* "co_js" is the list of sound problems whose unsound  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
816  | 
pendants couldn't be satisfied and hence that most  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
817  | 
probably can't be satisfied themselves. *)  | 
| 33192 | 818  | 
val co_js =  | 
819  | 
map (fn j => j - 1) unsat_js  | 
|
820  | 
|> filter (fn j =>  | 
|
821  | 
j >= 0 andalso  | 
|
822  | 
scopes_equivalent  | 
|
| 35814 | 823  | 
(#scope (snd (nth problems j)),  | 
824  | 
#scope (snd (nth problems (j + 1)))))  | 
|
| 33192 | 825  | 
val bye_js = sort_distinct int_ord (map fst sat_ps @  | 
826  | 
unsat_js @ co_js)  | 
|
827  | 
val problems =  | 
|
828  | 
problems |> filter_out_indices bye_js  | 
|
829  | 
|> max_potential <= 0  | 
|
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
830  | 
? filter_out (#unsound o snd)  | 
| 33192 | 831  | 
in  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
832  | 
solve_any_problem (found_really_genuine, max_potential,  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
833  | 
max_genuine, donno) false problems  | 
| 33192 | 834  | 
end  | 
835  | 
end  | 
|
836  | 
else  | 
|
837  | 
let  | 
|
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
838  | 
val genuine_codes =  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
839  | 
con_ps |> take max_genuine  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
840  | 
|> map (print_and_check true)  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
841  | 
val max_genuine = max_genuine - length genuine_codes  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
842  | 
val found_really_genuine =  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
843  | 
found_really_genuine orelse exists fst genuine_codes  | 
| 33192 | 844  | 
in  | 
845  | 
if max_genuine <= 0 orelse not first_time then  | 
|
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
846  | 
(found_really_genuine, 0, max_genuine, donno)  | 
| 33192 | 847  | 
else  | 
848  | 
let  | 
|
849  | 
val bye_js = sort_distinct int_ord  | 
|
850  | 
(map fst sat_ps @ unsat_js)  | 
|
851  | 
val problems =  | 
|
852  | 
problems |> filter_out_indices bye_js  | 
|
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
853  | 
|> filter_out (#unsound o snd)  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
854  | 
in  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
855  | 
solve_any_problem (found_really_genuine, 0, max_genuine,  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
856  | 
donno) false problems  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
857  | 
end  | 
| 33192 | 858  | 
end  | 
859  | 
end  | 
|
| 34126 | 860  | 
| KK.TimedOut unsat_js =>  | 
| 33192 | 861  | 
(update_checked_problems problems unsat_js; raise TimeLimit.TimeOut)  | 
| 34126 | 862  | 
| KK.Error (s, unsat_js) =>  | 
| 33192 | 863  | 
(update_checked_problems problems unsat_js;  | 
864  | 
             print_v (K ("Kodkod error: " ^ s ^ "."));
 | 
|
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
865  | 
(found_really_genuine, max_potential, max_genuine, donno + 1))  | 
| 33192 | 866  | 
end  | 
867  | 
||
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
868  | 
fun run_batch j n scopes (found_really_genuine, max_potential, max_genuine,  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
869  | 
donno) =  | 
| 33192 | 870  | 
let  | 
871  | 
val _ =  | 
|
872  | 
if null scopes then  | 
|
| 47560 | 873  | 
print_nt (K "The scope specification is inconsistent.")  | 
| 33192 | 874  | 
else if verbose then  | 
| 47560 | 875  | 
pprint_nt (fn () => Pretty.chunks  | 
| 33192 | 876  | 
[Pretty.blk (0,  | 
877  | 
pstrs ((if n > 1 then  | 
|
878  | 
"Batch " ^ string_of_int (j + 1) ^ " of " ^  | 
|
879  | 
signed_string_of_int n ^ ": "  | 
|
880  | 
else  | 
|
881  | 
"") ^  | 
|
882  | 
"Trying " ^ string_of_int (length scopes) ^  | 
|
883  | 
" scope" ^ plural_s_for_list scopes ^ ":")),  | 
|
884  | 
Pretty.indent indent_size  | 
|
885  | 
(Pretty.chunks (map2  | 
|
886  | 
(fn j => fn scope =>  | 
|
887  | 
Pretty.block (  | 
|
888  | 
(case pretties_for_scope scope true of  | 
|
889  | 
[] => [Pretty.str "Empty"]  | 
|
890  | 
| pretties => pretties) @  | 
|
891  | 
[Pretty.str (if j = 1 then "." else ";")]))  | 
|
892  | 
(length scopes downto 1) scopes))])  | 
|
893  | 
else  | 
|
894  | 
()  | 
|
| 
35280
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
35220 
diff
changeset
 | 
895  | 
fun add_problem_for_scope (scope, unsound) (problems, donno) =  | 
| 33192 | 896  | 
(check_deadline ();  | 
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
897  | 
case problem_for_scope unsound scope of  | 
| 33192 | 898  | 
SOME problem =>  | 
899  | 
(problems  | 
|
900  | 
|> (null problems orelse  | 
|
| 35814 | 901  | 
not (KK.problems_equivalent (fst problem, fst (hd problems))))  | 
| 33192 | 902  | 
? cons problem, donno)  | 
903  | 
| NONE => (problems, donno + 1))  | 
|
904  | 
val (problems, donno) =  | 
|
905  | 
fold add_problem_for_scope  | 
|
906  | 
(map_product pair scopes  | 
|
907  | 
((if max_genuine > 0 then [false] else []) @  | 
|
908  | 
(if max_potential > 0 then [true] else [])))  | 
|
909  | 
([], donno)  | 
|
910  | 
val _ = Unsynchronized.change generated_problems (append problems)  | 
|
911  | 
val _ = Unsynchronized.change generated_scopes (append scopes)  | 
|
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
912  | 
val _ =  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
913  | 
if j + 1 = n then  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
914  | 
let  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
915  | 
val (unsound_problems, sound_problems) =  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
916  | 
List.partition (#unsound o snd) (!generated_problems)  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
917  | 
in  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
918  | 
if not (null sound_problems) andalso  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
919  | 
forall (KK.is_problem_trivially_false o fst)  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
920  | 
sound_problems then  | 
| 47560 | 921  | 
print_nt (fn () =>  | 
| 44395 | 922  | 
"Warning: The conjecture " ^  | 
923  | 
(if falsify then "either trivially holds"  | 
|
924  | 
else "is either trivially unsatisfiable") ^  | 
|
925  | 
" for the given scopes or lies outside Nitpick's supported \  | 
|
| 35384 | 926  | 
\fragment." ^  | 
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
927  | 
(if exists (not o KK.is_problem_trivially_false o fst)  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
928  | 
unsound_problems then  | 
| 
41992
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
929  | 
" Only potentially spurious " ^ das_wort_model ^  | 
| 
 
0e4716fa330a
reword Nitpick's wording concerning potential counterexamples
 
blanchet 
parents: 
41985 
diff
changeset
 | 
930  | 
"s may be found."  | 
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
931  | 
else  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
932  | 
""))  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
933  | 
else  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
934  | 
()  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
935  | 
end  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
936  | 
else  | 
| 
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
937  | 
()  | 
| 33192 | 938  | 
in  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
939  | 
solve_any_problem (found_really_genuine, max_potential, max_genuine,  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
940  | 
donno) true (rev problems)  | 
| 33192 | 941  | 
end  | 
942  | 
||
943  | 
fun scope_count (problems : rich_problem list) scope =  | 
|
| 35814 | 944  | 
length (filter (curry scopes_equivalent scope o #scope o snd) problems)  | 
| 33192 | 945  | 
fun excipit did_so_and_so =  | 
946  | 
let  | 
|
947  | 
val do_filter =  | 
|
| 
35185
 
9b8f351cced6
added yet another hint to Nitpick's output, this time warning about problems for which nothing was effectively tested
 
blanchet 
parents: 
35183 
diff
changeset
 | 
948  | 
if !met_potential = max_potential then filter_out (#unsound o snd)  | 
| 33192 | 949  | 
else I  | 
950  | 
val total = length (!scopes)  | 
|
951  | 
val unsat =  | 
|
952  | 
fold (fn scope =>  | 
|
953  | 
case scope_count (do_filter (!generated_problems)) scope of  | 
|
954  | 
0 => I  | 
|
955  | 
| n =>  | 
|
| 
33556
 
cba22e2999d5
renamed Nitpick option "coalesce_type_vars" to "merge_type_vars" (shorter) and cleaned up old hacks that are no longer necessary
 
blanchet 
parents: 
33233 
diff
changeset
 | 
956  | 
scope_count (do_filter (these (!checked_problems)))  | 
| 
 
cba22e2999d5
renamed Nitpick option "coalesce_type_vars" to "merge_type_vars" (shorter) and cleaned up old hacks that are no longer necessary
 
blanchet 
parents: 
33233 
diff
changeset
 | 
957  | 
scope = n  | 
| 
 
cba22e2999d5
renamed Nitpick option "coalesce_type_vars" to "merge_type_vars" (shorter) and cleaned up old hacks that are no longer necessary
 
blanchet 
parents: 
33233 
diff
changeset
 | 
958  | 
? Integer.add 1) (!generated_scopes) 0  | 
| 33192 | 959  | 
in  | 
| 47560 | 960  | 
(if mode = TPTP then "% SZS status Unknown\n" else "") ^  | 
| 33192 | 961  | 
"Nitpick " ^ did_so_and_so ^  | 
962  | 
(if is_some (!checked_problems) andalso total > 0 then  | 
|
| 39361 | 963  | 
" " ^ string_of_int unsat ^ " of " ^ string_of_int total ^ " scope"  | 
964  | 
^ plural_s total  | 
|
| 33192 | 965  | 
else  | 
966  | 
"") ^ "."  | 
|
967  | 
end  | 
|
968  | 
||
| 
37704
 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 
blanchet 
parents: 
37497 
diff
changeset
 | 
969  | 
val (skipped, the_scopes) =  | 
| 
 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 
blanchet 
parents: 
37497 
diff
changeset
 | 
970  | 
all_scopes hol_ctxt binarize cards_assigns maxes_assigns iters_assigns  | 
| 
 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 
blanchet 
parents: 
37497 
diff
changeset
 | 
971  | 
bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs  | 
| 
 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 
blanchet 
parents: 
37497 
diff
changeset
 | 
972  | 
finitizable_dataTs  | 
| 
 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 
blanchet 
parents: 
37497 
diff
changeset
 | 
973  | 
val _ = if skipped > 0 then  | 
| 47560 | 974  | 
print_nt (fn () => "Too many scopes. Skipping " ^  | 
975  | 
string_of_int skipped ^ " scope" ^  | 
|
976  | 
plural_s skipped ^  | 
|
977  | 
". (Consider using \"mono\" or \  | 
|
978  | 
\\"merge_type_vars\" to prevent this.)")  | 
|
| 
37704
 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 
blanchet 
parents: 
37497 
diff
changeset
 | 
979  | 
else  | 
| 
 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 
blanchet 
parents: 
37497 
diff
changeset
 | 
980  | 
()  | 
| 
 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 
blanchet 
parents: 
37497 
diff
changeset
 | 
981  | 
val _ = scopes := the_scopes  | 
| 
 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 
blanchet 
parents: 
37497 
diff
changeset
 | 
982  | 
|
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
983  | 
fun run_batches _ _ []  | 
| 
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
984  | 
(found_really_genuine, max_potential, max_genuine, donno) =  | 
| 33192 | 985  | 
if donno > 0 andalso max_genuine > 0 then  | 
| 47560 | 986  | 
(print_nt (fn () => excipit "checked"); unknownN)  | 
| 33192 | 987  | 
else if max_genuine = original_max_genuine then  | 
988  | 
if max_potential = original_max_potential then  | 
|
| 47562 | 989  | 
(print_t (K "% SZS status Unknown");  | 
| 55888 | 990  | 
print_nt (fn () => "Nitpick found no " ^ das_wort_model ^ ".");  | 
991  | 
if skipped > 0 then unknownN else noneN)  | 
|
| 33192 | 992  | 
else  | 
| 47560 | 993  | 
(print_nt (fn () =>  | 
| 
38123
 
36f649db4a6c
clarify Nitpick's output in case of a potential counterexample
 
blanchet 
parents: 
37928 
diff
changeset
 | 
994  | 
                 excipit ("could not find a" ^
 | 
| 
 
36f649db4a6c
clarify Nitpick's output in case of a potential counterexample
 
blanchet 
parents: 
37928 
diff
changeset
 | 
995  | 
(if max_genuine = 1 then  | 
| 
 
36f649db4a6c
clarify Nitpick's output in case of a potential counterexample
 
blanchet 
parents: 
37928 
diff
changeset
 | 
996  | 
" better " ^ das_wort_model ^ "."  | 
| 
 
36f649db4a6c
clarify Nitpick's output in case of a potential counterexample
 
blanchet 
parents: 
37928 
diff
changeset
 | 
997  | 
else  | 
| 
 
36f649db4a6c
clarify Nitpick's output in case of a potential counterexample
 
blanchet 
parents: 
37928 
diff
changeset
 | 
998  | 
"ny better " ^ das_wort_model ^ "s.") ^  | 
| 
 
36f649db4a6c
clarify Nitpick's output in case of a potential counterexample
 
blanchet 
parents: 
37928 
diff
changeset
 | 
999  | 
" It checked"));  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
1000  | 
potentialN)  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
1001  | 
else if found_really_genuine then  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
1002  | 
genuineN  | 
| 33192 | 1003  | 
else  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
1004  | 
quasi_genuineN  | 
| 33192 | 1005  | 
| run_batches j n (batch :: batches) z =  | 
| 
35671
 
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
 
blanchet 
parents: 
35665 
diff
changeset
 | 
1006  | 
let val (z as (_, _, max_genuine, _)) = run_batch j n batch z in  | 
| 33192 | 1007  | 
run_batches (j + 1) n (if max_genuine > 0 then batches else []) z  | 
1008  | 
end  | 
|
1009  | 
||
| 48323 | 1010  | 
val batches = chunk_list batch_size (!scopes)  | 
| 33192 | 1011  | 
val outcome_code =  | 
| 
40411
 
36b7ed41ca9f
removed explicit "Interrupt" handling for conformity with async model -- unfortunately the user loses the information about how many scopes were checked, but this needs to be retought with the new interface anyway
 
blanchet 
parents: 
40381 
diff
changeset
 | 
1012  | 
run_batches 0 (length batches) batches  | 
| 
 
36b7ed41ca9f
removed explicit "Interrupt" handling for conformity with async model -- unfortunately the user loses the information about how many scopes were checked, but this needs to be retought with the new interface anyway
 
blanchet 
parents: 
40381 
diff
changeset
 | 
1013  | 
(false, max_potential, max_genuine, 0)  | 
| 33192 | 1014  | 
handle TimeLimit.TimeOut =>  | 
| 47560 | 1015  | 
(print_nt (fn () => excipit "ran out of time after checking");  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
1016  | 
if !met_potential > 0 then potentialN else unknownN)  | 
| 
40341
 
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
 
blanchet 
parents: 
40223 
diff
changeset
 | 
1017  | 
val _ = print_v (fn () =>  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
50830 
diff
changeset
 | 
1018  | 
"Total time: " ^ string_of_time (Timer.checkRealTimer timer) ^  | 
| 
40341
 
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
 
blanchet 
parents: 
40223 
diff
changeset
 | 
1019  | 
".")  | 
| 
53815
 
e8aa538e959e
encode goal digest in spying log (to detect duplicates)
 
blanchet 
parents: 
53802 
diff
changeset
 | 
1020  | 
val _ = spying spy (fn () => (state, i, "outcome: " ^ outcome_code))  | 
| 33192 | 1021  | 
in (outcome_code, !state_ref) end  | 
1022  | 
||
| 
41051
 
2ed1b971fc20
give the inner timeout mechanism a chance, since it gives more precise information to the user
 
blanchet 
parents: 
41048 
diff
changeset
 | 
1023  | 
(* Give the inner timeout a chance. *)  | 
| 
41772
 
27d4c768cf20
make Nitpick's timeout mechanism somewhat more reliable/friendly;
 
blanchet 
parents: 
41278 
diff
changeset
 | 
1024  | 
val timeout_bonus = seconds 1.0  | 
| 
41051
 
2ed1b971fc20
give the inner timeout mechanism a chance, since it gives more precise information to the user
 
blanchet 
parents: 
41048 
diff
changeset
 | 
1025  | 
|
| 
55539
 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 
blanchet 
parents: 
54816 
diff
changeset
 | 
1026  | 
fun pick_nits_in_term state (params as {timeout, expect, ...}) mode i n step
 | 
| 
 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 
blanchet 
parents: 
54816 
diff
changeset
 | 
1027  | 
subst def_assm_ts nondef_assm_ts orig_t =  | 
| 47670 | 1028  | 
let  | 
1029  | 
val print_nt = if is_mode_nt mode then Output.urgent_message else K ()  | 
|
1030  | 
val print_t = if mode = TPTP then Output.urgent_message else K ()  | 
|
1031  | 
in  | 
|
| 
37213
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1032  | 
if getenv "KODKODI" = "" then  | 
| 
37497
 
71fdbffe3275
distinguish between "unknown" and "no Kodkodi installed" errors
 
blanchet 
parents: 
37273 
diff
changeset
 | 
1033  | 
(* The "expect" argument is deliberately ignored if Kodkodi is missing so  | 
| 
 
71fdbffe3275
distinguish between "unknown" and "no Kodkodi installed" errors
 
blanchet 
parents: 
37273 
diff
changeset
 | 
1034  | 
that the "Nitpick_Examples" can be processed on any machine. *)  | 
| 55889 | 1035  | 
(print_nt (Pretty.string_of (plazy (K kodkodi_not_installed_message)));  | 
| 
37497
 
71fdbffe3275
distinguish between "unknown" and "no Kodkodi installed" errors
 
blanchet 
parents: 
37273 
diff
changeset
 | 
1036  | 
       ("no_kodkodi", state))
 | 
| 
37213
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1037  | 
else  | 
| 
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1038  | 
let  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
1039  | 
val unknown_outcome = (unknownN, state)  | 
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
1040  | 
val deadline = Time.+ (Time.now (), timeout)  | 
| 
37213
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1041  | 
val outcome as (outcome_code, _) =  | 
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
53815 
diff
changeset
 | 
1042  | 
TimeLimit.timeLimit (Time.+ (timeout, timeout_bonus))  | 
| 43022 | 1043  | 
(pick_them_nits_in_term deadline state params mode i n step subst  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
1044  | 
def_assm_ts nondef_assm_ts) orig_t  | 
| 
37213
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1045  | 
handle TOO_LARGE (_, details) =>  | 
| 47670 | 1046  | 
(print_t "% SZS status Unknown";  | 
1047  | 
                  print_nt ("Limit reached: " ^ details ^ "."); unknown_outcome)
 | 
|
| 
37213
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1048  | 
| TOO_SMALL (_, details) =>  | 
| 47670 | 1049  | 
(print_t "% SZS status Unknown";  | 
1050  | 
                  print_nt ("Limit reached: " ^ details ^ "."); unknown_outcome)
 | 
|
| 
37213
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1051  | 
| Kodkod.SYNTAX (_, details) =>  | 
| 47670 | 1052  | 
(print_t "% SZS status Unknown";  | 
1053  | 
                  print_nt ("Malformed Kodkodi output: " ^ details ^ ".");
 | 
|
| 
37213
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1054  | 
unknown_outcome)  | 
| 
41772
 
27d4c768cf20
make Nitpick's timeout mechanism somewhat more reliable/friendly;
 
blanchet 
parents: 
41278 
diff
changeset
 | 
1055  | 
| TimeLimit.TimeOut =>  | 
| 47670 | 1056  | 
(print_t "% SZS status TimedOut";  | 
1057  | 
print_nt "Nitpick ran out of time."; unknown_outcome)  | 
|
| 
37213
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1058  | 
in  | 
| 
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1059  | 
if expect = "" orelse outcome_code = expect then outcome  | 
| 
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1060  | 
        else error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
 | 
| 
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1061  | 
end  | 
| 
 
efcad7594872
fix handling of "split" w.r.t. new definition + fix exception handling w.r.t. "expect" option
 
blanchet 
parents: 
37146 
diff
changeset
 | 
1062  | 
end  | 
| 33192 | 1063  | 
|
| 42486 | 1064  | 
fun is_fixed_equation ctxt  | 
| 56245 | 1065  | 
                      (Const (@{const_name Pure.eq}, _) $ Free (s, _) $ Const _) =
 | 
| 42486 | 1066  | 
Variable.is_fixed ctxt s  | 
| 35335 | 1067  | 
| is_fixed_equation _ _ = false  | 
| 55889 | 1068  | 
|
| 35335 | 1069  | 
fun extract_fixed_frees ctxt (assms, t) =  | 
1070  | 
let  | 
|
1071  | 
val (subst, other_assms) =  | 
|
| 42486 | 1072  | 
List.partition (is_fixed_equation ctxt) assms  | 
| 35335 | 1073  | 
|>> map Logic.dest_equals  | 
1074  | 
in (subst, other_assms, subst_atomic subst t) end  | 
|
1075  | 
||
| 43022 | 1076  | 
fun pick_nits_in_subgoal state params mode i step =  | 
| 33192 | 1077  | 
let  | 
1078  | 
val ctxt = Proof.context_of state  | 
|
| 33292 | 1079  | 
val t = state |> Proof.raw_goal |> #goal |> prop_of  | 
| 33192 | 1080  | 
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: 
34938 
diff
changeset
 | 
1081  | 
case Logic.count_prems t of  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
42959 
diff
changeset
 | 
1082  | 
0 => (Output.urgent_message "No subgoal!"; (noneN, state))  | 
| 
34982
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34938 
diff
changeset
 | 
1083  | 
| n =>  | 
| 33192 | 1084  | 
let  | 
| 36406 | 1085  | 
val t = Logic.goal_params t i |> fst  | 
| 33192 | 1086  | 
val assms = map term_of (Assumption.all_assms_of ctxt)  | 
| 35335 | 1087  | 
val (subst, assms, t) = extract_fixed_frees ctxt (assms, t)  | 
| 
47715
 
04400144c6fc
handle TPTP definitions as definitions in Nitpick rather than as axioms
 
blanchet 
parents: 
47670 
diff
changeset
 | 
1088  | 
in pick_nits_in_term state params mode i n step subst [] assms t end  | 
| 33192 | 1089  | 
end  | 
1090  | 
||
1091  | 
end;  |