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