| author | kuncar | 
| Tue, 09 Oct 2012 16:58:36 +0200 | |
| changeset 49758 | 718f10c8bbfc | 
| parent 47909 | 5f1afeebafbc | 
| child 52174 | 7fd0b5cfbb79 | 
| permissions | -rw-r--r-- | 
| 33982 | 1  | 
(* Title: HOL/Tools/Nitpick/nitpick_scope.ML  | 
| 33192 | 2  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
34982
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34936 
diff
changeset
 | 
3  | 
Copyright 2008, 2009, 2010  | 
| 33192 | 4  | 
|
5  | 
Scope enumerator for Nitpick.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature NITPICK_SCOPE =  | 
|
9  | 
sig  | 
|
| 
33705
 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 
blanchet 
parents: 
33580 
diff
changeset
 | 
10  | 
type styp = Nitpick_Util.styp  | 
| 
35070
 
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
 
blanchet 
parents: 
34982 
diff
changeset
 | 
11  | 
type hol_context = Nitpick_HOL.hol_context  | 
| 33192 | 12  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
13  | 
type constr_spec =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
14  | 
    {const: styp,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
15  | 
delta: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
16  | 
epsilon: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
17  | 
exclusive: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
18  | 
explicit_max: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
19  | 
total: bool}  | 
| 33192 | 20  | 
|
| 38126 | 21  | 
type datatype_spec =  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
22  | 
    {typ: typ,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
23  | 
card: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
24  | 
co: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
25  | 
standard: bool,  | 
| 38126 | 26  | 
self_rec: bool,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
27  | 
complete: bool * bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
28  | 
concrete: bool * bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
29  | 
deep: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
30  | 
constrs: constr_spec list}  | 
| 33192 | 31  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
32  | 
type scope =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
33  | 
    {hol_ctxt: hol_context,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
34  | 
binarize: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
35  | 
card_assigns: (typ * int) list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
36  | 
bits: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
37  | 
bisim_depth: int,  | 
| 38126 | 38  | 
datatypes: datatype_spec list,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
39  | 
ofs: int Typtab.table}  | 
| 33192 | 40  | 
|
| 38127 | 41  | 
val is_asymmetric_nondatatype : typ -> bool  | 
| 38126 | 42  | 
val datatype_spec : datatype_spec list -> typ -> datatype_spec option  | 
43  | 
val constr_spec : datatype_spec list -> styp -> constr_spec  | 
|
44  | 
val is_complete_type : datatype_spec list -> bool -> typ -> bool  | 
|
45  | 
val is_concrete_type : datatype_spec list -> bool -> typ -> bool  | 
|
46  | 
val is_exact_type : datatype_spec list -> bool -> typ -> bool  | 
|
| 33192 | 47  | 
val offset_of_type : int Typtab.table -> typ -> int  | 
48  | 
val spec_of_type : scope -> typ -> int * int  | 
|
49  | 
val pretties_for_scope : scope -> bool -> Pretty.T list  | 
|
50  | 
val multiline_string_for_scope : scope -> string  | 
|
| 35814 | 51  | 
val scopes_equivalent : scope * scope -> bool  | 
| 33192 | 52  | 
val scope_less_eq : scope -> scope -> bool  | 
| 38126 | 53  | 
val is_self_recursive_constr_type : typ -> bool  | 
| 33192 | 54  | 
val all_scopes :  | 
| 
36386
 
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
 
blanchet 
parents: 
36385 
diff
changeset
 | 
55  | 
hol_context -> bool -> (typ option * int list) list  | 
| 33192 | 56  | 
-> (styp option * int list) list -> (styp option * int list) list  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
57  | 
-> int list -> int list -> typ list -> typ list -> typ list -> typ list  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
58  | 
-> int * scope list  | 
| 33192 | 59  | 
end;  | 
60  | 
||
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
61  | 
structure Nitpick_Scope : NITPICK_SCOPE =  | 
| 33192 | 62  | 
struct  | 
63  | 
||
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
64  | 
open Nitpick_Util  | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
65  | 
open Nitpick_HOL  | 
| 33192 | 66  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
67  | 
type constr_spec =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
68  | 
  {const: styp,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
69  | 
delta: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
70  | 
epsilon: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
71  | 
exclusive: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
72  | 
explicit_max: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
73  | 
total: bool}  | 
| 33192 | 74  | 
|
| 38126 | 75  | 
type datatype_spec =  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
76  | 
  {typ: typ,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
77  | 
card: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
78  | 
co: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
79  | 
standard: bool,  | 
| 38126 | 80  | 
self_rec: bool,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
81  | 
complete: bool * bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
82  | 
concrete: bool * bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
83  | 
deep: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
84  | 
constrs: constr_spec list}  | 
| 33192 | 85  | 
|
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
86  | 
type scope =  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
87  | 
  {hol_ctxt: hol_context,
 | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
88  | 
binarize: bool,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
89  | 
card_assigns: (typ * int) list,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
90  | 
bits: int,  | 
| 
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
91  | 
bisim_depth: int,  | 
| 38126 | 92  | 
datatypes: datatype_spec list,  | 
| 
36390
 
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
 
blanchet 
parents: 
36386 
diff
changeset
 | 
93  | 
ofs: int Typtab.table}  | 
| 33192 | 94  | 
|
95  | 
datatype row_kind = Card of typ | Max of styp  | 
|
96  | 
||
97  | 
type row = row_kind * int list  | 
|
98  | 
type block = row list  | 
|
99  | 
||
| 38127 | 100  | 
val is_asymmetric_nondatatype =  | 
101  | 
is_iterator_type orf is_integer_type orf is_bit_type  | 
|
102  | 
||
| 38126 | 103  | 
fun datatype_spec (dtypes : datatype_spec list) T =  | 
| 
34121
 
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
 
blanchet 
parents: 
33982 
diff
changeset
 | 
104  | 
List.find (curry (op =) T o #typ) dtypes  | 
| 33192 | 105  | 
|
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
106  | 
fun constr_spec [] x = raise TERM ("Nitpick_Scope.constr_spec", [Const x])
 | 
| 38126 | 107  | 
  | constr_spec ({constrs, ...} :: dtypes : datatype_spec list) (x as (s, T)) =
 | 
| 
34121
 
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
 
blanchet 
parents: 
33982 
diff
changeset
 | 
108  | 
case List.find (curry (op =) (s, body_type T) o (apsnd body_type o #const))  | 
| 33192 | 109  | 
constrs of  | 
110  | 
SOME c => c  | 
|
111  | 
| NONE => constr_spec dtypes x  | 
|
112  | 
||
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
113  | 
fun is_complete_type dtypes facto (Type (@{type_name fun}, [T1, T2])) =
 | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
114  | 
is_concrete_type dtypes facto T1 andalso is_complete_type dtypes facto T2  | 
| 
38190
 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 
blanchet 
parents: 
38188 
diff
changeset
 | 
115  | 
  | is_complete_type dtypes facto (Type (@{type_name prod}, Ts)) =
 | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
116  | 
forall (is_complete_type dtypes facto) Ts  | 
| 46083 | 117  | 
  | is_complete_type dtypes facto (Type (@{type_name set}, [T'])) =
 | 
118  | 
is_concrete_type dtypes facto T'  | 
|
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
119  | 
| is_complete_type dtypes facto T =  | 
| 
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
 | 
120  | 
not (is_integer_like_type T) andalso not (is_bit_type T) andalso  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
121  | 
fun_from_pair (#complete (the (datatype_spec dtypes T))) facto  | 
| 33192 | 122  | 
handle Option.Option => true  | 
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
123  | 
and is_concrete_type dtypes facto (Type (@{type_name fun}, [T1, T2])) =
 | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
124  | 
is_complete_type dtypes facto T1 andalso is_concrete_type dtypes facto T2  | 
| 
38190
 
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
 
blanchet 
parents: 
38188 
diff
changeset
 | 
125  | 
  | is_concrete_type dtypes facto (Type (@{type_name prod}, Ts)) =
 | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
126  | 
forall (is_concrete_type dtypes facto) Ts  | 
| 46083 | 127  | 
  | is_concrete_type dtypes facto (Type (@{type_name set}, [T'])) =
 | 
128  | 
is_complete_type dtypes facto T'  | 
|
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
129  | 
| is_concrete_type dtypes facto T =  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
130  | 
fun_from_pair (#concrete (the (datatype_spec dtypes T))) facto  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
131  | 
handle Option.Option => true  | 
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
132  | 
and is_exact_type dtypes facto =  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
133  | 
is_complete_type dtypes facto andf is_concrete_type dtypes facto  | 
| 33192 | 134  | 
|
135  | 
fun offset_of_type ofs T =  | 
|
136  | 
case Typtab.lookup ofs T of  | 
|
137  | 
SOME j0 => j0  | 
|
138  | 
| NONE => Typtab.lookup ofs dummyT |> the_default 0  | 
|
139  | 
||
140  | 
fun spec_of_type ({card_assigns, ofs, ...} : scope) T =
 | 
|
141  | 
(card_of_type card_assigns T  | 
|
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
142  | 
   handle TYPE ("Nitpick_HOL.card_of_type", _, _) => ~1, offset_of_type ofs T)
 | 
| 33192 | 143  | 
|
| 38188 | 144  | 
fun quintuple_for_scope code_type code_term code_string  | 
| 
39118
 
12f3788be67b
turned show_all_types into proper configuration option;
 
wenzelm 
parents: 
38240 
diff
changeset
 | 
145  | 
        ({hol_ctxt = {ctxt = ctxt0, stds, ...}, card_assigns, bits, bisim_depth,
 | 
| 
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
 | 
146  | 
datatypes, ...} : scope) =  | 
| 33192 | 147  | 
let  | 
| 
39118
 
12f3788be67b
turned show_all_types into proper configuration option;
 
wenzelm 
parents: 
38240 
diff
changeset
 | 
148  | 
val ctxt = set_show_all_types ctxt0  | 
| 
35075
 
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
 
blanchet 
parents: 
35072 
diff
changeset
 | 
149  | 
    val boring_Ts = [@{typ unsigned_bit}, @{typ signed_bit},
 | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
150  | 
                     @{typ bisim_iterator}]
 | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
151  | 
val (iter_assigns, card_assigns) =  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
152  | 
card_assigns |> filter_out (member (op =) boring_Ts o fst)  | 
| 33192 | 153  | 
|> List.partition (is_fp_iterator_type o fst)  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
154  | 
val (secondary_card_assigns, primary_card_assigns) =  | 
| 
37256
 
0dca1ec52999
thread along context instead of theory for typedef lookup
 
blanchet 
parents: 
36390 
diff
changeset
 | 
155  | 
card_assigns  | 
| 
 
0dca1ec52999
thread along context instead of theory for typedef lookup
 
blanchet 
parents: 
36390 
diff
changeset
 | 
156  | 
|> List.partition ((is_integer_type orf is_datatype ctxt stds) o fst)  | 
| 33192 | 157  | 
val cards =  | 
| 38188 | 158  | 
map (fn (T, k) =>  | 
159  | 
              [code_type ctxt T, code_string (" = " ^ string_of_int k)])
 | 
|
| 33192 | 160  | 
fun maxes () =  | 
161  | 
maps (map_filter  | 
|
162  | 
                (fn {const, explicit_max, ...} =>
 | 
|
163  | 
if explicit_max < 0 then  | 
|
164  | 
NONE  | 
|
165  | 
else  | 
|
| 38188 | 166  | 
SOME [code_term ctxt (Const const),  | 
167  | 
                            code_string (" = " ^ string_of_int explicit_max)])
 | 
|
| 33192 | 168  | 
o #constrs) datatypes  | 
169  | 
fun iters () =  | 
|
170  | 
map (fn (T, k) =>  | 
|
| 38188 | 171  | 
[code_term ctxt (Const (const_for_iterator_type T)),  | 
172  | 
               code_string (" = " ^ string_of_int (k - 1))]) iter_assigns
 | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
173  | 
fun miscs () =  | 
| 38188 | 174  | 
(if bits = 0 then []  | 
175  | 
       else [code_string ("bits = " ^ string_of_int bits)]) @
 | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
176  | 
(if bisim_depth < 0 andalso forall (not o #co) datatypes then []  | 
| 38188 | 177  | 
       else [code_string ("bisim_depth = " ^ signed_string_of_int bisim_depth)])
 | 
| 33192 | 178  | 
in  | 
| 
39118
 
12f3788be67b
turned show_all_types into proper configuration option;
 
wenzelm 
parents: 
38240 
diff
changeset
 | 
179  | 
(cards primary_card_assigns, cards secondary_card_assigns,  | 
| 
 
12f3788be67b
turned show_all_types into proper configuration option;
 
wenzelm 
parents: 
38240 
diff
changeset
 | 
180  | 
maxes (), iters (), miscs ())  | 
| 33192 | 181  | 
end  | 
182  | 
||
183  | 
fun pretties_for_scope scope verbose =  | 
|
184  | 
let  | 
|
| 38188 | 185  | 
fun standard_blocks s = map (Pretty.block o cons (Pretty.str (s ^ " ")))  | 
186  | 
val (primary_cards, secondary_cards, maxes, iters, miscs) =  | 
|
187  | 
quintuple_for_scope (pretty_maybe_quote oo pretty_for_type)  | 
|
188  | 
(pretty_maybe_quote oo Syntax.pretty_term)  | 
|
189  | 
Pretty.str scope  | 
|
| 33192 | 190  | 
in  | 
| 38188 | 191  | 
standard_blocks "card" primary_cards @  | 
192  | 
(if verbose then  | 
|
193  | 
standard_blocks "card" secondary_cards @  | 
|
194  | 
standard_blocks "max" maxes @  | 
|
195  | 
standard_blocks "iter" iters @  | 
|
196  | 
miscs  | 
|
197  | 
else  | 
|
198  | 
[])  | 
|
199  | 
|> pretty_serial_commas "and"  | 
|
| 33192 | 200  | 
end  | 
201  | 
||
202  | 
fun multiline_string_for_scope scope =  | 
|
203  | 
let  | 
|
| 38188 | 204  | 
val (primary_cards, secondary_cards, maxes, iters, miscs) =  | 
205  | 
quintuple_for_scope Syntax.string_of_typ Syntax.string_of_term I scope  | 
|
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
206  | 
val cards = primary_cards @ secondary_cards  | 
| 33192 | 207  | 
in  | 
| 38188 | 208  | 
case (if null cards then [] else ["card: " ^ commas (map implode cards)]) @  | 
209  | 
(if null maxes then [] else ["max: " ^ commas (map implode maxes)]) @  | 
|
210  | 
(if null iters then [] else ["iter: " ^ commas (map implode iters)]) @  | 
|
211  | 
miscs of  | 
|
| 33192 | 212  | 
[] => "empty"  | 
213  | 
| lines => space_implode "\n" lines  | 
|
214  | 
end  | 
|
215  | 
||
| 35814 | 216  | 
fun scopes_equivalent (s1 : scope, s2 : scope) =  | 
| 33192 | 217  | 
#datatypes s1 = #datatypes s2 andalso #card_assigns s1 = #card_assigns s2  | 
218  | 
fun scope_less_eq (s1 : scope) (s2 : scope) =  | 
|
219  | 
(s1, s2) |> pairself (map snd o #card_assigns) |> op ~~ |> forall (op <=)  | 
|
220  | 
||
221  | 
fun rank_of_row (_, ks) = length ks  | 
|
222  | 
fun rank_of_block block = fold Integer.max (map rank_of_row block) 1  | 
|
| 
41991
 
ea02b9ee3085
prevent an exception if "card" is empty (e.g., "nitpick [card]")
 
blanchet 
parents: 
41052 
diff
changeset
 | 
223  | 
fun project_row _ (y, []) = (y, [1]) (* desperate measure *)  | 
| 
 
ea02b9ee3085
prevent an exception if "card" is empty (e.g., "nitpick [card]")
 
blanchet 
parents: 
41052 
diff
changeset
 | 
224  | 
| project_row column (y, ks) = (y, [nth ks (Int.min (column, length ks - 1))])  | 
| 33192 | 225  | 
fun project_block (column, block) = map (project_row column) block  | 
226  | 
||
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
227  | 
fun lookup_ints_assign eq assigns key =  | 
| 
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
228  | 
case triple_lookup eq assigns key of  | 
| 33192 | 229  | 
SOME ks => ks  | 
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
230  | 
  | NONE => raise ARG ("Nitpick_Scope.lookup_ints_assign", "")
 | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
231  | 
fun lookup_type_ints_assign thy assigns T =  | 
| 36384 | 232  | 
map (Integer.max 1) (lookup_ints_assign (type_match thy) assigns T)  | 
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
233  | 
  handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
 | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
234  | 
         raise TYPE ("Nitpick_Scope.lookup_type_ints_assign", [T], [])
 | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
235  | 
fun lookup_const_ints_assign thy assigns x =  | 
| 
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
236  | 
lookup_ints_assign (const_match thy) assigns x  | 
| 
33232
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
237  | 
  handle ARG ("Nitpick_Scope.lookup_ints_assign", _) =>
 | 
| 
 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 
blanchet 
parents: 
33192 
diff
changeset
 | 
238  | 
         raise TERM ("Nitpick_Scope.lookup_const_ints_assign", [Const x])
 | 
| 33192 | 239  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
240  | 
fun row_for_constr thy maxes_assigns constr =  | 
| 
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
241  | 
SOME (Max constr, lookup_const_ints_assign thy maxes_assigns constr)  | 
| 33192 | 242  | 
  handle TERM ("lookup_const_ints_assign", _) => NONE
 | 
243  | 
||
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
244  | 
val max_bits = 31 (* Kodkod limit *)  | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
245  | 
|
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
246  | 
fun block_for_type (hol_ctxt as {thy, ...}) binarize cards_assigns maxes_assigns
 | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
247  | 
iters_assigns bitss bisim_depths T =  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
248  | 
case T of  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
249  | 
    @{typ unsigned_bit} =>
 | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
250  | 
[(Card T, map (Integer.min max_bits o Integer.max 1) bitss)]  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
251  | 
  | @{typ signed_bit} =>
 | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
252  | 
[(Card T, map (Integer.add 1 o Integer.min max_bits o Integer.max 1) bitss)]  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
253  | 
  | @{typ "unsigned_bit word"} =>
 | 
| 34126 | 254  | 
[(Card T, lookup_type_ints_assign thy cards_assigns nat_T)]  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
255  | 
  | @{typ "signed_bit word"} =>
 | 
| 34126 | 256  | 
[(Card T, lookup_type_ints_assign thy cards_assigns int_T)]  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
257  | 
  | @{typ bisim_iterator} =>
 | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
258  | 
[(Card T, map (Integer.add 1 o Integer.max 0) bisim_depths)]  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
259  | 
| _ =>  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
260  | 
if is_fp_iterator_type T then  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
261  | 
[(Card T, map (Integer.add 1 o Integer.max 0)  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
262  | 
(lookup_const_ints_assign thy iters_assigns  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
263  | 
(const_for_iterator_type T)))]  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
264  | 
else  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
265  | 
(Card T, lookup_type_ints_assign thy cards_assigns T) ::  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
266  | 
(case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
267  | 
[_] => []  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
268  | 
| constrs => map_filter (row_for_constr thy maxes_assigns) constrs)  | 
| 33192 | 269  | 
|
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
270  | 
fun blocks_for_types hol_ctxt binarize cards_assigns maxes_assigns iters_assigns  | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
271  | 
bitss bisim_depths mono_Ts nonmono_Ts =  | 
| 33192 | 272  | 
let  | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
273  | 
val block_for = block_for_type hol_ctxt binarize cards_assigns maxes_assigns  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
274  | 
iters_assigns bitss bisim_depths  | 
| 33192 | 275  | 
val mono_block = maps block_for mono_Ts  | 
276  | 
val nonmono_blocks = map block_for nonmono_Ts  | 
|
277  | 
in mono_block :: nonmono_blocks end  | 
|
278  | 
||
279  | 
val sync_threshold = 5  | 
|
| 
38186
 
c28018f5a1d6
example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
 
blanchet 
parents: 
38180 
diff
changeset
 | 
280  | 
val linearity = 5  | 
| 33192 | 281  | 
|
| 
38186
 
c28018f5a1d6
example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
 
blanchet 
parents: 
38180 
diff
changeset
 | 
282  | 
val all_combinations_ordered_smartly =  | 
| 33192 | 283  | 
let  | 
| 
38186
 
c28018f5a1d6
example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
 
blanchet 
parents: 
38180 
diff
changeset
 | 
284  | 
fun cost [] = 0  | 
| 
 
c28018f5a1d6
example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
 
blanchet 
parents: 
38180 
diff
changeset
 | 
285  | 
| cost [k] = k  | 
| 
 
c28018f5a1d6
example tweaking -- also prevents Nitpick_Tests from using more than 1 thread
 
blanchet 
parents: 
38180 
diff
changeset
 | 
286  | 
| cost (k :: ks) =  | 
| 
34121
 
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
 
blanchet 
parents: 
33982 
diff
changeset
 | 
287  | 
if k < sync_threshold andalso forall (curry (op =) k) ks then  | 
| 33192 | 288  | 
k - sync_threshold  | 
289  | 
else  | 
|
| 38187 | 290  | 
k :: ks |> map (fn k => (k + linearity) * (k + linearity))  | 
291  | 
|> Integer.sum  | 
|
| 33192 | 292  | 
in  | 
| 38187 | 293  | 
all_combinations #> map (`cost) #> sort (int_ord o pairself fst) #> map snd  | 
| 33192 | 294  | 
end  | 
295  | 
||
296  | 
fun is_self_recursive_constr_type T =  | 
|
| 
34121
 
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
 
blanchet 
parents: 
33982 
diff
changeset
 | 
297  | 
exists (exists_subtype (curry (op =) (body_type T))) (binder_types T)  | 
| 33192 | 298  | 
|
299  | 
fun constr_max maxes x = the_default ~1 (AList.lookup (op =) maxes x)  | 
|
300  | 
||
301  | 
type scope_desc = (typ * int) list * (styp * int) list  | 
|
302  | 
||
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
303  | 
fun is_surely_inconsistent_card_assign hol_ctxt binarize  | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
304  | 
(card_assigns, max_assigns) (T, k) =  | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
305  | 
case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of  | 
| 33192 | 306  | 
[] => false  | 
307  | 
| xs =>  | 
|
308  | 
let  | 
|
| 34126 | 309  | 
val dom_cards =  | 
310  | 
map (Integer.prod o map (bounded_card_of_type k ~1 card_assigns)  | 
|
| 33192 | 311  | 
o binder_types o snd) xs  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
312  | 
val maxes = map (constr_max max_assigns) xs  | 
| 34126 | 313  | 
fun effective_max card ~1 = card  | 
| 33192 | 314  | 
| effective_max card max = Int.min (card, max)  | 
| 34126 | 315  | 
val max = map2 effective_max dom_cards maxes |> Integer.sum  | 
316  | 
in max < k end  | 
|
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
317  | 
fun is_surely_inconsistent_scope_description hol_ctxt binarize seen rest  | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
318  | 
max_assigns =  | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
319  | 
exists (is_surely_inconsistent_card_assign hol_ctxt binarize  | 
| 34126 | 320  | 
(seen @ rest, max_assigns)) seen  | 
| 33192 | 321  | 
|
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
322  | 
fun repair_card_assigns hol_ctxt binarize (card_assigns, max_assigns) =  | 
| 33192 | 323  | 
let  | 
324  | 
fun aux seen [] = SOME seen  | 
|
| 
35280
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
35220 
diff
changeset
 | 
325  | 
| aux _ ((_, 0) :: _) = NONE  | 
| 34126 | 326  | 
| aux seen ((T, k) :: rest) =  | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
327  | 
(if is_surely_inconsistent_scope_description hol_ctxt binarize  | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
328  | 
((T, k) :: seen) rest max_assigns then  | 
| 33192 | 329  | 
raise SAME ()  | 
330  | 
else  | 
|
| 34126 | 331  | 
case aux ((T, k) :: seen) rest of  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
332  | 
SOME assigns => SOME assigns  | 
| 33192 | 333  | 
| NONE => raise SAME ())  | 
| 34126 | 334  | 
handle SAME () => aux seen ((T, k - 1) :: rest)  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
335  | 
in aux [] (rev card_assigns) end  | 
| 33192 | 336  | 
|
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
337  | 
fun repair_iterator_assign ctxt assigns (T as Type (_, Ts), k) =  | 
| 33192 | 338  | 
    (T, if T = @{typ bisim_iterator} then
 | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
339  | 
let  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
340  | 
val co_cards = map snd (filter (is_codatatype ctxt o fst) assigns)  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
341  | 
in Int.min (k, Integer.sum co_cards) end  | 
| 33192 | 342  | 
else if is_fp_iterator_type T then  | 
343  | 
case Ts of  | 
|
344  | 
[] => 1  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
345  | 
| _ => bounded_card_of_type k ~1 assigns (foldr1 HOLogic.mk_prodT Ts)  | 
| 33192 | 346  | 
else  | 
347  | 
k)  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
348  | 
| repair_iterator_assign _ _ assign = assign  | 
| 33192 | 349  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
350  | 
fun add_row_to_scope_descriptor (kind, ks) (card_assigns, max_assigns) =  | 
| 33192 | 351  | 
case kind of  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
352  | 
Card T => ((T, the_single ks) :: card_assigns, max_assigns)  | 
| 
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
353  | 
| Max x => (card_assigns, (x, the_single ks) :: max_assigns)  | 
| 33192 | 354  | 
fun scope_descriptor_from_block block =  | 
355  | 
fold_rev add_row_to_scope_descriptor block ([], [])  | 
|
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
356  | 
fun scope_descriptor_from_combination (hol_ctxt as {ctxt, ...}) binarize blocks
 | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
357  | 
columns =  | 
| 33192 | 358  | 
let  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
359  | 
val (card_assigns, max_assigns) =  | 
| 33192 | 360  | 
maps project_block (columns ~~ blocks) |> scope_descriptor_from_block  | 
361  | 
in  | 
|
| 
41991
 
ea02b9ee3085
prevent an exception if "card" is empty (e.g., "nitpick [card]")
 
blanchet 
parents: 
41052 
diff
changeset
 | 
362  | 
(card_assigns, max_assigns)  | 
| 
 
ea02b9ee3085
prevent an exception if "card" is empty (e.g., "nitpick [card]")
 
blanchet 
parents: 
41052 
diff
changeset
 | 
363  | 
|> repair_card_assigns hol_ctxt binarize  | 
| 
 
ea02b9ee3085
prevent an exception if "card" is empty (e.g., "nitpick [card]")
 
blanchet 
parents: 
41052 
diff
changeset
 | 
364  | 
|> Option.map  | 
| 
 
ea02b9ee3085
prevent an exception if "card" is empty (e.g., "nitpick [card]")
 
blanchet 
parents: 
41052 
diff
changeset
 | 
365  | 
(fn card_assigns =>  | 
| 
 
ea02b9ee3085
prevent an exception if "card" is empty (e.g., "nitpick [card]")
 
blanchet 
parents: 
41052 
diff
changeset
 | 
366  | 
(map (repair_iterator_assign ctxt card_assigns) card_assigns,  | 
| 
 
ea02b9ee3085
prevent an exception if "card" is empty (e.g., "nitpick [card]")
 
blanchet 
parents: 
41052 
diff
changeset
 | 
367  | 
max_assigns))  | 
| 33192 | 368  | 
end  | 
369  | 
||
| 38124 | 370  | 
fun offset_table_for_card_assigns dtypes assigns =  | 
| 33192 | 371  | 
let  | 
372  | 
fun aux next _ [] = Typtab.update_new (dummyT, next)  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
373  | 
| aux next reusable ((T, k) :: assigns) =  | 
| 38127 | 374  | 
if k = 1 orelse is_asymmetric_nondatatype T then  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
375  | 
aux next reusable assigns  | 
| 33192 | 376  | 
else if length (these (Option.map #constrs (datatype_spec dtypes T)))  | 
377  | 
> 1 then  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
378  | 
Typtab.update_new (T, next) #> aux (next + k) reusable assigns  | 
| 33192 | 379  | 
else  | 
380  | 
case AList.lookup (op =) reusable k of  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
381  | 
SOME j0 => Typtab.update_new (T, j0) #> aux next reusable assigns  | 
| 33192 | 382  | 
| NONE => Typtab.update_new (T, next)  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
383  | 
#> aux (next + k) ((k, next) :: reusable) assigns  | 
| 38124 | 384  | 
in Typtab.empty |> aux 0 [] assigns end  | 
| 33192 | 385  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
386  | 
fun domain_card max card_assigns =  | 
| 
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
387  | 
Integer.prod o map (bounded_card_of_type max max card_assigns) o binder_types  | 
| 33192 | 388  | 
|
| 
38162
 
824e940a3dd0
minor symmetry breaking for codatatypes like llist
 
blanchet 
parents: 
38127 
diff
changeset
 | 
389  | 
fun add_constr_spec (card_assigns, max_assigns) acyclic card sum_dom_cards  | 
| 
35280
 
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
 
blanchet 
parents: 
35220 
diff
changeset
 | 
390  | 
num_self_recs num_non_self_recs (self_rec, x as (_, T))  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
391  | 
constrs =  | 
| 33192 | 392  | 
let  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
393  | 
val max = constr_max max_assigns x  | 
| 33192 | 394  | 
fun next_delta () = if null constrs then 0 else #epsilon (hd constrs)  | 
395  | 
    val {delta, epsilon, exclusive, total} =
 | 
|
396  | 
if max = 0 then  | 
|
397  | 
let val delta = next_delta () in  | 
|
398  | 
          {delta = delta, epsilon = delta, exclusive = true, total = false}
 | 
|
399  | 
end  | 
|
| 
38162
 
824e940a3dd0
minor symmetry breaking for codatatypes like llist
 
blanchet 
parents: 
38127 
diff
changeset
 | 
400  | 
else if num_self_recs > 0 then  | 
| 38193 | 401  | 
(if num_non_self_recs = 1 then  | 
| 
35072
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
402  | 
if self_rec then  | 
| 38193 | 403  | 
case List.last constrs of  | 
404  | 
               {delta = 0, epsilon = 1, exclusive = true, ...} =>
 | 
|
405  | 
               {delta = 1, epsilon = card, exclusive = (num_self_recs = 1),
 | 
|
406  | 
total = false}  | 
|
| 
35072
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
407  | 
| _ => raise SAME ()  | 
| 
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
408  | 
else  | 
| 
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
409  | 
if domain_card 2 card_assigns T = 1 then  | 
| 
38162
 
824e940a3dd0
minor symmetry breaking for codatatypes like llist
 
blanchet 
parents: 
38127 
diff
changeset
 | 
410  | 
               {delta = 0, epsilon = 1, exclusive = acyclic, total = acyclic}
 | 
| 
35072
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
411  | 
else  | 
| 
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
412  | 
raise SAME ()  | 
| 
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
413  | 
else  | 
| 
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
414  | 
raise SAME ())  | 
| 
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
415  | 
handle SAME () =>  | 
| 
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
416  | 
               {delta = 0, epsilon = card, exclusive = false, total = false}
 | 
| 33192 | 417  | 
else if card = sum_dom_cards (card + 1) then  | 
418  | 
let val delta = next_delta () in  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
419  | 
          {delta = delta, epsilon = delta + domain_card card card_assigns T,
 | 
| 33192 | 420  | 
exclusive = true, total = true}  | 
421  | 
end  | 
|
422  | 
else  | 
|
423  | 
        {delta = 0, epsilon = card,
 | 
|
424  | 
exclusive = (num_self_recs + num_non_self_recs = 1), total = false}  | 
|
425  | 
in  | 
|
426  | 
    {const = x, delta = delta, epsilon = epsilon, exclusive = exclusive,
 | 
|
427  | 
explicit_max = max, total = total} :: constrs  | 
|
428  | 
end  | 
|
429  | 
||
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
430  | 
fun has_exact_card hol_ctxt facto finitizable_dataTs card_assigns T =  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
431  | 
let val card = card_of_type card_assigns T in  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
432  | 
card = bounded_exact_card_of_type hol_ctxt  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
433  | 
(if facto then finitizable_dataTs else []) (card + 1) 0  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
434  | 
card_assigns T  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
435  | 
end  | 
| 
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
436  | 
|
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
437  | 
fun datatype_spec_from_scope_descriptor (hol_ctxt as {thy, ctxt, stds, ...})
 | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
438  | 
binarize deep_dataTs finitizable_dataTs (desc as (card_assigns, _))  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
439  | 
(T, card) =  | 
| 33192 | 440  | 
let  | 
| 
34982
 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 
blanchet 
parents: 
34936 
diff
changeset
 | 
441  | 
val deep = member (op =) deep_dataTs T  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
442  | 
val co = is_codatatype ctxt T  | 
| 
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
 | 
443  | 
val standard = is_standard_datatype thy stds T  | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
444  | 
val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T  | 
| 33192 | 445  | 
val self_recs = map (is_self_recursive_constr_type o snd) xs  | 
446  | 
val (num_self_recs, num_non_self_recs) =  | 
|
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
447  | 
List.partition I self_recs |> pairself length  | 
| 38126 | 448  | 
val self_rec = num_self_recs > 0  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
449  | 
fun is_complete facto =  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
450  | 
has_exact_card hol_ctxt facto finitizable_dataTs card_assigns T  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
451  | 
fun is_concrete facto =  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
452  | 
is_word_type T orelse  | 
| 
47909
 
5f1afeebafbc
fixed "real" after they were redefined as a 'quotient_type'
 
blanchet 
parents: 
46083 
diff
changeset
 | 
453  | 
(* FIXME: looks wrong; other types than just functions might be  | 
| 45402 | 454  | 
abstract. "is_complete" is also suspicious. *)  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
455  | 
xs |> maps (binder_types o snd) |> maps binder_types  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
456  | 
|> forall (has_exact_card hol_ctxt facto finitizable_dataTs  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
457  | 
card_assigns)  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
458  | 
val complete = pair_from_fun is_complete  | 
| 
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
459  | 
val concrete = pair_from_fun is_concrete  | 
| 33192 | 460  | 
fun sum_dom_cards max =  | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
461  | 
map (domain_card max card_assigns o snd) xs |> Integer.sum  | 
| 33192 | 462  | 
val constrs =  | 
| 
38162
 
824e940a3dd0
minor symmetry breaking for codatatypes like llist
 
blanchet 
parents: 
38127 
diff
changeset
 | 
463  | 
fold_rev (add_constr_spec desc (not co andalso standard) card  | 
| 
 
824e940a3dd0
minor symmetry breaking for codatatypes like llist
 
blanchet 
parents: 
38127 
diff
changeset
 | 
464  | 
sum_dom_cards num_self_recs num_non_self_recs)  | 
| 
35072
 
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
 
blanchet 
parents: 
35071 
diff
changeset
 | 
465  | 
(sort (bool_ord o swap o pairself fst) (self_recs ~~ xs)) []  | 
| 
33558
 
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
 
blanchet 
parents: 
33232 
diff
changeset
 | 
466  | 
in  | 
| 38126 | 467  | 
    {typ = T, card = card, co = co, standard = standard, self_rec = self_rec,
 | 
468  | 
complete = complete, concrete = concrete, deep = deep, constrs = constrs}  | 
|
| 
33558
 
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
 
blanchet 
parents: 
33232 
diff
changeset
 | 
469  | 
end  | 
| 33192 | 470  | 
|
| 
37256
 
0dca1ec52999
thread along context instead of theory for typedef lookup
 
blanchet 
parents: 
36390 
diff
changeset
 | 
471  | 
fun scope_from_descriptor (hol_ctxt as {ctxt, stds, ...}) binarize deep_dataTs
 | 
| 
36386
 
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
 
blanchet 
parents: 
36385 
diff
changeset
 | 
472  | 
finitizable_dataTs (desc as (card_assigns, _)) =  | 
| 33192 | 473  | 
let  | 
| 
33558
 
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
 
blanchet 
parents: 
33232 
diff
changeset
 | 
474  | 
val datatypes =  | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
475  | 
map (datatype_spec_from_scope_descriptor hol_ctxt binarize deep_dataTs  | 
| 
35385
 
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
 
blanchet 
parents: 
35280 
diff
changeset
 | 
476  | 
finitizable_dataTs desc)  | 
| 
37256
 
0dca1ec52999
thread along context instead of theory for typedef lookup
 
blanchet 
parents: 
36390 
diff
changeset
 | 
477  | 
(filter (is_datatype ctxt stds o fst) card_assigns)  | 
| 
34124
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
478  | 
    val bits = card_of_type card_assigns @{typ signed_bit} - 1
 | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
479  | 
               handle TYPE ("Nitpick_HOL.card_of_type", _, _) =>
 | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
480  | 
                      card_of_type card_assigns @{typ unsigned_bit}
 | 
| 
 
c4628a1dcf75
added support for binary nat/int representation to Nitpick
 
blanchet 
parents: 
34123 
diff
changeset
 | 
481  | 
                      handle TYPE ("Nitpick_HOL.card_of_type", _, _) => 0
 | 
| 
34123
 
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
 
blanchet 
parents: 
34121 
diff
changeset
 | 
482  | 
    val bisim_depth = card_of_type card_assigns @{typ bisim_iterator} - 1
 | 
| 33192 | 483  | 
in  | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
484  | 
    {hol_ctxt = hol_ctxt, binarize = binarize, card_assigns = card_assigns,
 | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
485  | 
datatypes = datatypes, bits = bits, bisim_depth = bisim_depth,  | 
| 38124 | 486  | 
ofs = offset_table_for_card_assigns datatypes card_assigns}  | 
| 33192 | 487  | 
end  | 
488  | 
||
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
489  | 
fun repair_cards_assigns_wrt_boxing_etc _ _ [] = []  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
490  | 
| repair_cards_assigns_wrt_boxing_etc thy Ts ((SOME T, ks) :: cards_assigns) =  | 
| 33192 | 491  | 
(if is_fun_type T orelse is_pair_type T then  | 
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
492  | 
Ts |> filter (curry (type_match thy o swap) T) |> map (rpair ks o SOME)  | 
| 33192 | 493  | 
else  | 
| 
35665
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
494  | 
[(SOME T, ks)]) @  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
495  | 
repair_cards_assigns_wrt_boxing_etc thy Ts cards_assigns  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
496  | 
| repair_cards_assigns_wrt_boxing_etc thy Ts ((NONE, ks) :: cards_assigns) =  | 
| 
 
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
 
blanchet 
parents: 
35385 
diff
changeset
 | 
497  | 
(NONE, ks) :: repair_cards_assigns_wrt_boxing_etc thy Ts cards_assigns  | 
| 33192 | 498  | 
|
| 
38180
 
7a88032f9265
bump up the max cardinalities, to use up more of the time given to us by the user
 
blanchet 
parents: 
38162 
diff
changeset
 | 
499  | 
val max_scopes = 5000  | 
| 
 
7a88032f9265
bump up the max cardinalities, to use up more of the time given to us by the user
 
blanchet 
parents: 
38162 
diff
changeset
 | 
500  | 
val distinct_threshold = 1000  | 
| 33192 | 501  | 
|
| 
36386
 
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
 
blanchet 
parents: 
36385 
diff
changeset
 | 
502  | 
fun all_scopes (hol_ctxt as {thy, ...}) binarize cards_assigns maxes_assigns
 | 
| 
 
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
 
blanchet 
parents: 
36385 
diff
changeset
 | 
503  | 
iters_assigns bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs  | 
| 
 
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
 
blanchet 
parents: 
36385 
diff
changeset
 | 
504  | 
finitizable_dataTs =  | 
| 33192 | 505  | 
let  | 
| 
38240
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
506  | 
val cards_assigns =  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
507  | 
repair_cards_assigns_wrt_boxing_etc thy mono_Ts cards_assigns  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
508  | 
val blocks =  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
509  | 
blocks_for_types hol_ctxt binarize cards_assigns maxes_assigns  | 
| 
 
a44d108a8d39
local versions of Nitpick.register_xxx functions
 
blanchet 
parents: 
38193 
diff
changeset
 | 
510  | 
iters_assigns bitss bisim_depths mono_Ts nonmono_Ts  | 
| 33192 | 511  | 
val ranks = map rank_of_block blocks  | 
| 
33580
 
45c33e97cb86
added datatype constructor cache in Nitpick (to speed up the scope enumeration) and never test more than 4096 scopes
 
blanchet 
parents: 
33558 
diff
changeset
 | 
512  | 
val all = all_combinations_ordered_smartly (map (rpair 0) ranks)  | 
| 33957 | 513  | 
val head = take max_scopes all  | 
| 
35190
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
514  | 
val descs =  | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
515  | 
map_filter (scope_descriptor_from_combination hol_ctxt binarize blocks)  | 
| 
 
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
 
blanchet 
parents: 
35179 
diff
changeset
 | 
516  | 
head  | 
| 33192 | 517  | 
in  | 
| 
33580
 
45c33e97cb86
added datatype constructor cache in Nitpick (to speed up the scope enumeration) and never test more than 4096 scopes
 
blanchet 
parents: 
33558 
diff
changeset
 | 
518  | 
(length all - length head,  | 
| 
 
45c33e97cb86
added datatype constructor cache in Nitpick (to speed up the scope enumeration) and never test more than 4096 scopes
 
blanchet 
parents: 
33558 
diff
changeset
 | 
519  | 
descs |> length descs <= distinct_threshold ? distinct (op =)  | 
| 
36386
 
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
 
blanchet 
parents: 
36385 
diff
changeset
 | 
520  | 
|> map (scope_from_descriptor hol_ctxt binarize deep_dataTs  | 
| 
 
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
 
blanchet 
parents: 
36385 
diff
changeset
 | 
521  | 
finitizable_dataTs))  | 
| 33192 | 522  | 
end  | 
523  | 
||
524  | 
end;  |