author | haftmann |
Wed, 14 Jul 2010 14:20:47 +0200 | |
changeset 37819 | 000049335247 |
parent 37678 | 0040bafffdef |
child 38124 | 6538e25cf5dd |
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 |
|
36390
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
21 |
type dtype_spec = |
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, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
26 |
complete: bool * bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
27 |
concrete: bool * bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
28 |
deep: bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
29 |
constrs: constr_spec list} |
33192 | 30 |
|
36390
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
31 |
type scope = |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
32 |
{hol_ctxt: hol_context, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
33 |
binarize: bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
34 |
card_assigns: (typ * int) list, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
35 |
bits: int, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
36 |
bisim_depth: int, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
37 |
datatypes: dtype_spec list, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
38 |
ofs: int Typtab.table} |
33192 | 39 |
|
40 |
val datatype_spec : dtype_spec list -> typ -> dtype_spec option |
|
41 |
val constr_spec : dtype_spec list -> styp -> constr_spec |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
42 |
val is_complete_type : dtype_spec list -> bool -> typ -> bool |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
43 |
val is_concrete_type : dtype_spec list -> bool -> typ -> bool |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
44 |
val is_exact_type : dtype_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 |
51 |
val all_scopes : |
|
36386
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents:
36385
diff
changeset
|
52 |
hol_context -> bool -> (typ option * int list) list |
33192 | 53 |
-> (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
|
54 |
-> 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
|
55 |
-> int * scope list |
33192 | 56 |
end; |
57 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
58 |
structure Nitpick_Scope : NITPICK_SCOPE = |
33192 | 59 |
struct |
60 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
61 |
open Nitpick_Util |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
62 |
open Nitpick_HOL |
33192 | 63 |
|
36390
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
64 |
type constr_spec = |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
65 |
{const: styp, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
66 |
delta: int, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
67 |
epsilon: int, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
68 |
exclusive: bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
69 |
explicit_max: int, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
70 |
total: bool} |
33192 | 71 |
|
36390
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
72 |
type dtype_spec = |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
73 |
{typ: typ, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
74 |
card: int, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
75 |
co: bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
76 |
standard: bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
77 |
complete: bool * bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
78 |
concrete: bool * bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
79 |
deep: bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
80 |
constrs: constr_spec list} |
33192 | 81 |
|
36390
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
82 |
type scope = |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
83 |
{hol_ctxt: hol_context, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
84 |
binarize: bool, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
85 |
card_assigns: (typ * int) list, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
86 |
bits: int, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
87 |
bisim_depth: int, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
88 |
datatypes: dtype_spec list, |
eee4ee6a5cbe
remove "show_skolems" option and change style of record declarations
blanchet
parents:
36386
diff
changeset
|
89 |
ofs: int Typtab.table} |
33192 | 90 |
|
91 |
datatype row_kind = Card of typ | Max of styp |
|
92 |
||
93 |
type row = row_kind * int list |
|
94 |
type block = row list |
|
95 |
||
96 |
fun datatype_spec (dtypes : dtype_spec list) T = |
|
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
97 |
List.find (curry (op =) T o #typ) dtypes |
33192 | 98 |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
99 |
fun constr_spec [] x = raise TERM ("Nitpick_Scope.constr_spec", [Const x]) |
33192 | 100 |
| constr_spec ({constrs, ...} :: dtypes : dtype_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
|
101 |
case List.find (curry (op =) (s, body_type T) o (apsnd body_type o #const)) |
33192 | 102 |
constrs of |
103 |
SOME c => c |
|
104 |
| NONE => constr_spec dtypes x |
|
105 |
||
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
106 |
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
|
107 |
is_concrete_type dtypes facto T1 andalso is_complete_type dtypes facto T2 |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
108 |
| is_complete_type dtypes facto (Type (@{type_name fin_fun}, [T1, T2])) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
109 |
is_exact_type dtypes facto T1 andalso is_complete_type dtypes facto T2 |
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
37256
diff
changeset
|
110 |
| is_complete_type dtypes facto (Type (@{type_name Product_Type.prod}, Ts)) = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
111 |
forall (is_complete_type dtypes facto) Ts |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
112 |
| 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
|
113 |
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
|
114 |
fun_from_pair (#complete (the (datatype_spec dtypes T))) facto |
33192 | 115 |
handle Option.Option => true |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
116 |
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
|
117 |
is_complete_type dtypes facto T1 andalso is_concrete_type dtypes facto T2 |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
118 |
| is_concrete_type dtypes facto (Type (@{type_name fin_fun}, [_, T2])) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
119 |
is_concrete_type dtypes facto T2 |
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
37256
diff
changeset
|
120 |
| is_concrete_type dtypes facto (Type (@{type_name Product_Type.prod}, Ts)) = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
121 |
forall (is_concrete_type dtypes facto) Ts |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
122 |
| is_concrete_type dtypes facto T = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
123 |
fun_from_pair (#concrete (the (datatype_spec dtypes T))) facto |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
124 |
handle Option.Option => true |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
125 |
and is_exact_type dtypes facto = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
126 |
is_complete_type dtypes facto andf is_concrete_type dtypes facto |
33192 | 127 |
|
128 |
fun offset_of_type ofs T = |
|
129 |
case Typtab.lookup ofs T of |
|
130 |
SOME j0 => j0 |
|
131 |
| NONE => Typtab.lookup ofs dummyT |> the_default 0 |
|
132 |
||
133 |
fun spec_of_type ({card_assigns, ofs, ...} : scope) T = |
|
134 |
(card_of_type card_assigns T |
|
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
135 |
handle TYPE ("Nitpick_HOL.card_of_type", _, _) => ~1, offset_of_type ofs T) |
33192 | 136 |
|
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
|
137 |
fun quintuple_for_scope quote |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36390
diff
changeset
|
138 |
({hol_ctxt = {ctxt, 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
|
139 |
datatypes, ...} : scope) = |
33192 | 140 |
let |
35075
888802be2019
handle Nitpick's nonstandard model enumeration in a cleaner way;
blanchet
parents:
35072
diff
changeset
|
141 |
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
|
142 |
@{typ bisim_iterator}] |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
143 |
val (iter_assigns, card_assigns) = |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
144 |
card_assigns |> filter_out (member (op =) boring_Ts o fst) |
33192 | 145 |
|> List.partition (is_fp_iterator_type o fst) |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
146 |
val (secondary_card_assigns, primary_card_assigns) = |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36390
diff
changeset
|
147 |
card_assigns |
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36390
diff
changeset
|
148 |
|> List.partition ((is_integer_type orf is_datatype ctxt stds) o fst) |
33192 | 149 |
val cards = |
150 |
map (fn (T, k) => quote (string_for_type ctxt T) ^ " = " ^ |
|
151 |
string_of_int k) |
|
152 |
fun maxes () = |
|
153 |
maps (map_filter |
|
154 |
(fn {const, explicit_max, ...} => |
|
155 |
if explicit_max < 0 then |
|
156 |
NONE |
|
157 |
else |
|
158 |
SOME (Syntax.string_of_term ctxt (Const const) ^ " = " ^ |
|
159 |
string_of_int explicit_max)) |
|
160 |
o #constrs) datatypes |
|
161 |
fun iters () = |
|
162 |
map (fn (T, k) => |
|
163 |
quote (Syntax.string_of_term ctxt |
|
164 |
(Const (const_for_iterator_type T))) ^ " = " ^ |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
165 |
string_of_int (k - 1)) iter_assigns |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
166 |
fun miscs () = |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
167 |
(if bits = 0 then [] else ["bits = " ^ string_of_int bits]) @ |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
168 |
(if bisim_depth < 0 andalso forall (not o #co) datatypes then [] |
35178 | 169 |
else ["bisim_depth = " ^ signed_string_of_int bisim_depth]) |
33192 | 170 |
in |
171 |
setmp_show_all_types |
|
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
172 |
(fn () => (cards primary_card_assigns, cards secondary_card_assigns, |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
173 |
maxes (), iters (), miscs ())) () |
33192 | 174 |
end |
175 |
||
176 |
fun pretties_for_scope scope verbose = |
|
177 |
let |
|
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
178 |
val (primary_cards, secondary_cards, maxes, iters, bisim_depths) = |
33192 | 179 |
quintuple_for_scope maybe_quote scope |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
180 |
val ss = map (prefix "card ") primary_cards @ |
33192 | 181 |
(if verbose then |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
182 |
map (prefix "card ") secondary_cards @ |
33192 | 183 |
map (prefix "max ") maxes @ |
184 |
map (prefix "iter ") iters @ |
|
185 |
bisim_depths |
|
186 |
else |
|
187 |
[]) |
|
188 |
in |
|
36380
1e8fcaccb3e8
stop referring to Sledgehammer_Util stuff all over Nitpick code; instead, redeclare any needed function in Nitpick_Util as synonym for the Sledgehammer_Util function of the same name
blanchet
parents:
35968
diff
changeset
|
189 |
if null ss then [] |
1e8fcaccb3e8
stop referring to Sledgehammer_Util stuff all over Nitpick code; instead, redeclare any needed function in Nitpick_Util as synonym for the Sledgehammer_Util function of the same name
blanchet
parents:
35968
diff
changeset
|
190 |
else serial_commas "and" ss |> map Pretty.str |> Pretty.breaks |
33192 | 191 |
end |
192 |
||
193 |
fun multiline_string_for_scope scope = |
|
194 |
let |
|
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
195 |
val (primary_cards, secondary_cards, maxes, iters, bisim_depths) = |
33192 | 196 |
quintuple_for_scope I scope |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
197 |
val cards = primary_cards @ secondary_cards |
33192 | 198 |
in |
199 |
case (if null cards then [] else ["card: " ^ commas cards]) @ |
|
200 |
(if null maxes then [] else ["max: " ^ commas maxes]) @ |
|
201 |
(if null iters then [] else ["iter: " ^ commas iters]) @ |
|
202 |
bisim_depths of |
|
203 |
[] => "empty" |
|
204 |
| lines => space_implode "\n" lines |
|
205 |
end |
|
206 |
||
35814 | 207 |
fun scopes_equivalent (s1 : scope, s2 : scope) = |
33192 | 208 |
#datatypes s1 = #datatypes s2 andalso #card_assigns s1 = #card_assigns s2 |
209 |
fun scope_less_eq (s1 : scope) (s2 : scope) = |
|
210 |
(s1, s2) |> pairself (map snd o #card_assigns) |> op ~~ |> forall (op <=) |
|
211 |
||
212 |
fun rank_of_row (_, ks) = length ks |
|
213 |
fun rank_of_block block = fold Integer.max (map rank_of_row block) 1 |
|
214 |
fun project_row column (y, ks) = (y, [nth ks (Int.min (column, length ks - 1))]) |
|
215 |
fun project_block (column, block) = map (project_row column) block |
|
216 |
||
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
217 |
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
|
218 |
case triple_lookup eq assigns key of |
33192 | 219 |
SOME ks => ks |
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
220 |
| 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
|
221 |
fun lookup_type_ints_assign thy assigns T = |
36384 | 222 |
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
|
223 |
handle ARG ("Nitpick_Scope.lookup_ints_assign", _) => |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
224 |
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
|
225 |
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
|
226 |
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
|
227 |
handle ARG ("Nitpick_Scope.lookup_ints_assign", _) => |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
228 |
raise TERM ("Nitpick_Scope.lookup_const_ints_assign", [Const x]) |
33192 | 229 |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
230 |
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
|
231 |
SOME (Max constr, lookup_const_ints_assign thy maxes_assigns constr) |
33192 | 232 |
handle TERM ("lookup_const_ints_assign", _) => NONE |
233 |
||
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
234 |
val max_bits = 31 (* Kodkod limit *) |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
235 |
|
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
236 |
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
|
237 |
iters_assigns bitss bisim_depths T = |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
238 |
if T = @{typ unsigned_bit} then |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
239 |
[(Card T, map (Integer.min max_bits o Integer.max 1) bitss)] |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
240 |
else if T = @{typ signed_bit} then |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
241 |
[(Card T, map (Integer.add 1 o Integer.min max_bits o Integer.max 1) bitss)] |
34126 | 242 |
else if T = @{typ "unsigned_bit word"} then |
243 |
[(Card T, lookup_type_ints_assign thy cards_assigns nat_T)] |
|
244 |
else if T = @{typ "signed_bit word"} then |
|
245 |
[(Card T, lookup_type_ints_assign thy cards_assigns int_T)] |
|
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
246 |
else if T = @{typ bisim_iterator} then |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
247 |
[(Card T, map (Integer.add 1 o Integer.max 0) bisim_depths)] |
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
|
248 |
else if is_fp_iterator_type T then |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
249 |
[(Card T, map (Integer.add 1 o Integer.max 0) |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
250 |
(lookup_const_ints_assign thy iters_assigns |
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
|
251 |
(const_for_iterator_type T)))] |
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
|
252 |
else |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
253 |
(Card T, lookup_type_ints_assign thy cards_assigns T) :: |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
254 |
(case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of |
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
|
255 |
[_] => [] |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
256 |
| constrs => map_filter (row_for_constr thy maxes_assigns) constrs) |
33192 | 257 |
|
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
258 |
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
|
259 |
bitss bisim_depths mono_Ts nonmono_Ts = |
33192 | 260 |
let |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
261 |
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
|
262 |
iters_assigns bitss bisim_depths |
33192 | 263 |
val mono_block = maps block_for mono_Ts |
264 |
val nonmono_blocks = map block_for nonmono_Ts |
|
265 |
in mono_block :: nonmono_blocks end |
|
266 |
||
267 |
val sync_threshold = 5 |
|
268 |
||
269 |
fun all_combinations_ordered_smartly ks = |
|
270 |
let |
|
271 |
fun cost_with_monos [] = 0 |
|
272 |
| cost_with_monos (k :: ks) = |
|
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
273 |
if k < sync_threshold andalso forall (curry (op =) k) ks then |
33192 | 274 |
k - sync_threshold |
275 |
else |
|
276 |
k * (k + 1) div 2 + Integer.sum ks |
|
277 |
fun cost_without_monos [] = 0 |
|
278 |
| cost_without_monos [k] = k |
|
279 |
| cost_without_monos (_ :: k :: ks) = |
|
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
280 |
if k < sync_threshold andalso forall (curry (op =) k) ks then |
33192 | 281 |
k - sync_threshold |
282 |
else |
|
283 |
Integer.sum (k :: ks) |
|
284 |
in |
|
285 |
ks |> all_combinations |
|
286 |
|> map (`(if fst (hd ks) > 1 then cost_with_monos |
|
287 |
else cost_without_monos)) |
|
288 |
|> sort (int_ord o pairself fst) |> map snd |
|
289 |
end |
|
290 |
||
291 |
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
|
292 |
exists (exists_subtype (curry (op =) (body_type T))) (binder_types T) |
33192 | 293 |
|
294 |
fun constr_max maxes x = the_default ~1 (AList.lookup (op =) maxes x) |
|
295 |
||
296 |
type scope_desc = (typ * int) list * (styp * int) list |
|
297 |
||
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
298 |
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
|
299 |
(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
|
300 |
case binarized_and_boxed_datatype_constrs hol_ctxt binarize T of |
33192 | 301 |
[] => false |
302 |
| xs => |
|
303 |
let |
|
34126 | 304 |
val dom_cards = |
305 |
map (Integer.prod o map (bounded_card_of_type k ~1 card_assigns) |
|
33192 | 306 |
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
|
307 |
val maxes = map (constr_max max_assigns) xs |
34126 | 308 |
fun effective_max card ~1 = card |
33192 | 309 |
| effective_max card max = Int.min (card, max) |
34126 | 310 |
val max = map2 effective_max dom_cards maxes |> Integer.sum |
311 |
in max < k end |
|
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
312 |
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
|
313 |
max_assigns = |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
314 |
exists (is_surely_inconsistent_card_assign hol_ctxt binarize |
34126 | 315 |
(seen @ rest, max_assigns)) seen |
33192 | 316 |
|
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
317 |
fun repair_card_assigns hol_ctxt binarize (card_assigns, max_assigns) = |
33192 | 318 |
let |
319 |
fun aux seen [] = SOME seen |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
320 |
| aux _ ((_, 0) :: _) = NONE |
34126 | 321 |
| 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
|
322 |
(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
|
323 |
((T, k) :: seen) rest max_assigns then |
33192 | 324 |
raise SAME () |
325 |
else |
|
34126 | 326 |
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
|
327 |
SOME assigns => SOME assigns |
33192 | 328 |
| NONE => raise SAME ()) |
34126 | 329 |
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
|
330 |
in aux [] (rev card_assigns) end |
33192 | 331 |
|
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
332 |
fun repair_iterator_assign thy assigns (T as Type (_, Ts), k) = |
33192 | 333 |
(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
|
334 |
let |
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
335 |
val co_cards = map snd (filter (is_codatatype thy o fst) assigns) |
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
336 |
in Int.min (k, Integer.sum co_cards) end |
33192 | 337 |
else if is_fp_iterator_type T then |
338 |
case Ts of |
|
339 |
[] => 1 |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
340 |
| _ => bounded_card_of_type k ~1 assigns (foldr1 HOLogic.mk_prodT Ts) |
33192 | 341 |
else |
342 |
k) |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
343 |
| repair_iterator_assign _ _ assign = assign |
33192 | 344 |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
345 |
fun add_row_to_scope_descriptor (kind, ks) (card_assigns, max_assigns) = |
33192 | 346 |
case kind of |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
347 |
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
|
348 |
| Max x => (card_assigns, (x, the_single ks) :: max_assigns) |
33192 | 349 |
fun scope_descriptor_from_block block = |
350 |
fold_rev add_row_to_scope_descriptor block ([], []) |
|
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
351 |
fun scope_descriptor_from_combination (hol_ctxt as {thy, ...}) binarize blocks |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
352 |
columns = |
33192 | 353 |
let |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
354 |
val (card_assigns, max_assigns) = |
33192 | 355 |
maps project_block (columns ~~ blocks) |> scope_descriptor_from_block |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
356 |
val card_assigns = |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
357 |
repair_card_assigns hol_ctxt binarize (card_assigns, max_assigns) |> the |
33192 | 358 |
in |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
359 |
SOME (map (repair_iterator_assign thy card_assigns) card_assigns, |
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
360 |
max_assigns) |
33192 | 361 |
end |
362 |
handle Option.Option => NONE |
|
363 |
||
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
364 |
fun offset_table_for_card_assigns assigns dtypes = |
33192 | 365 |
let |
366 |
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
|
367 |
| aux next reusable ((T, k) :: assigns) = |
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
368 |
if k = 1 orelse is_iterator_type T orelse is_integer_type T |
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
369 |
orelse is_bit_type T then |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
370 |
aux next reusable assigns |
33192 | 371 |
else if length (these (Option.map #constrs (datatype_spec dtypes T))) |
372 |
> 1 then |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
373 |
Typtab.update_new (T, next) #> aux (next + k) reusable assigns |
33192 | 374 |
else |
375 |
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
|
376 |
SOME j0 => Typtab.update_new (T, j0) #> aux next reusable assigns |
33192 | 377 |
| 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
|
378 |
#> aux (next + k) ((k, next) :: reusable) assigns |
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
379 |
in aux 0 [] assigns Typtab.empty end |
33192 | 380 |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
381 |
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
|
382 |
Integer.prod o map (bounded_card_of_type max max card_assigns) o binder_types |
33192 | 383 |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
384 |
fun add_constr_spec (card_assigns, max_assigns) co card sum_dom_cards |
35280
54ab4921f826
fixed a few bugs in Nitpick and removed unreferenced variables
blanchet
parents:
35220
diff
changeset
|
385 |
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
|
386 |
constrs = |
33192 | 387 |
let |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
388 |
val max = constr_max max_assigns x |
33192 | 389 |
fun next_delta () = if null constrs then 0 else #epsilon (hd constrs) |
390 |
val {delta, epsilon, exclusive, total} = |
|
391 |
if max = 0 then |
|
392 |
let val delta = next_delta () in |
|
393 |
{delta = delta, epsilon = delta, exclusive = true, total = false} |
|
394 |
end |
|
395 |
else if not co andalso num_self_recs > 0 then |
|
35072
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
396 |
(if num_self_recs = 1 andalso num_non_self_recs = 1 then |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
397 |
if self_rec then |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
398 |
case constrs of |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
399 |
[{delta = 0, epsilon = 1, exclusive = true, ...}] => |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
400 |
{delta = 1, epsilon = card, exclusive = true, total = false} |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
401 |
| _ => raise SAME () |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
402 |
else |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
403 |
if domain_card 2 card_assigns T = 1 then |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
404 |
{delta = 0, epsilon = 1, exclusive = true, total = true} |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
405 |
else |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
406 |
raise SAME () |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
407 |
else |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
408 |
raise SAME ()) |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
409 |
handle SAME () => |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
410 |
{delta = 0, epsilon = card, exclusive = false, total = false} |
33192 | 411 |
else if card = sum_dom_cards (card + 1) then |
412 |
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
|
413 |
{delta = delta, epsilon = delta + domain_card card card_assigns T, |
33192 | 414 |
exclusive = true, total = true} |
415 |
end |
|
416 |
else |
|
417 |
{delta = 0, epsilon = card, |
|
418 |
exclusive = (num_self_recs + num_non_self_recs = 1), total = false} |
|
419 |
in |
|
420 |
{const = x, delta = delta, epsilon = epsilon, exclusive = exclusive, |
|
421 |
explicit_max = max, total = total} :: constrs |
|
422 |
end |
|
423 |
||
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
424 |
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
|
425 |
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
|
426 |
card = bounded_exact_card_of_type hol_ctxt |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
427 |
(if facto then finitizable_dataTs else []) (card + 1) 0 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
428 |
card_assigns T |
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
429 |
end |
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
430 |
|
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
|
431 |
fun datatype_spec_from_scope_descriptor (hol_ctxt as {thy, stds, ...}) binarize |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
432 |
deep_dataTs finitizable_dataTs (desc as (card_assigns, _)) (T, card) = |
33192 | 433 |
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
|
434 |
val deep = member (op =) deep_dataTs T |
33192 | 435 |
val co = is_codatatype thy 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
|
436 |
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
|
437 |
val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T |
33192 | 438 |
val self_recs = map (is_self_recursive_constr_type o snd) xs |
439 |
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
|
440 |
List.partition I self_recs |> pairself length |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
441 |
fun is_complete facto = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
442 |
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
|
443 |
fun is_concrete facto = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
444 |
is_word_type T orelse |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
445 |
xs |> maps (binder_types o snd) |> maps binder_types |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
446 |
|> forall (has_exact_card hol_ctxt facto finitizable_dataTs |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
447 |
card_assigns) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
448 |
val complete = pair_from_fun is_complete |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35280
diff
changeset
|
449 |
val concrete = pair_from_fun is_concrete |
33192 | 450 |
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
|
451 |
map (domain_card max card_assigns o snd) xs |> Integer.sum |
33192 | 452 |
val constrs = |
453 |
fold_rev (add_constr_spec desc co card sum_dom_cards num_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
|
454 |
num_non_self_recs) |
d79308423aea
optimize Nitpick's encoding for other datatypes than list that have a constant constructor like "Nil";
blanchet
parents:
35071
diff
changeset
|
455 |
(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
|
456 |
in |
35179
4b198af5beb5
redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents:
35178
diff
changeset
|
457 |
{typ = T, card = card, co = co, standard = standard, complete = complete, |
4b198af5beb5
redo Nitpick's nonstandard values as cyclic values (instead of additional constructors)
blanchet
parents:
35178
diff
changeset
|
458 |
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
|
459 |
end |
33192 | 460 |
|
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36390
diff
changeset
|
461 |
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
|
462 |
finitizable_dataTs (desc as (card_assigns, _)) = |
33192 | 463 |
let |
33558
a2db56854b83
optimized Nitpick's encoding and rendering of datatypes whose constructors don't appear in the problem
blanchet
parents:
33232
diff
changeset
|
464 |
val datatypes = |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
465 |
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
|
466 |
finitizable_dataTs desc) |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36390
diff
changeset
|
467 |
(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
|
468 |
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
|
469 |
handle TYPE ("Nitpick_HOL.card_of_type", _, _) => |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
470 |
card_of_type card_assigns @{typ unsigned_bit} |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
471 |
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
|
472 |
val bisim_depth = card_of_type card_assigns @{typ bisim_iterator} - 1 |
33192 | 473 |
in |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
474 |
{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
|
475 |
datatypes = datatypes, bits = bits, bisim_depth = bisim_depth, |
36386
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents:
36385
diff
changeset
|
476 |
ofs = offset_table_for_card_assigns card_assigns datatypes} |
33192 | 477 |
end |
478 |
||
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
479 |
fun repair_cards_assigns_wrt_boxing_etc _ _ [] = [] |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
480 |
| repair_cards_assigns_wrt_boxing_etc thy Ts ((SOME T, ks) :: cards_assigns) = |
33192 | 481 |
(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
|
482 |
Ts |> filter (curry (type_match thy o swap) T) |> map (rpair ks o SOME) |
33192 | 483 |
else |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
484 |
[(SOME T, ks)]) @ |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
485 |
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
|
486 |
| 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
|
487 |
(NONE, ks) :: repair_cards_assigns_wrt_boxing_etc thy Ts cards_assigns |
33192 | 488 |
|
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
|
489 |
val max_scopes = 4096 |
33192 | 490 |
val distinct_threshold = 512 |
491 |
||
36386
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents:
36385
diff
changeset
|
492 |
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
|
493 |
iters_assigns bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs |
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents:
36385
diff
changeset
|
494 |
finitizable_dataTs = |
33192 | 495 |
let |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
496 |
val cards_assigns = repair_cards_assigns_wrt_boxing_etc thy mono_Ts |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35385
diff
changeset
|
497 |
cards_assigns |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
498 |
val blocks = blocks_for_types hol_ctxt binarize cards_assigns maxes_assigns |
34124
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
499 |
iters_assigns bitss bisim_depths mono_Ts |
c4628a1dcf75
added support for binary nat/int representation to Nitpick
blanchet
parents:
34123
diff
changeset
|
500 |
nonmono_Ts |
33192 | 501 |
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
|
502 |
val all = all_combinations_ordered_smartly (map (rpair 0) ranks) |
33957 | 503 |
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
|
504 |
val descs = |
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35179
diff
changeset
|
505 |
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
|
506 |
head |
33192 | 507 |
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
|
508 |
(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
|
509 |
descs |> length descs <= distinct_threshold ? distinct (op =) |
36386
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents:
36385
diff
changeset
|
510 |
|> map (scope_from_descriptor hol_ctxt binarize deep_dataTs |
2132f15b366f
Fruhjahrsputz: remove three mostly useless Nitpick options
blanchet
parents:
36385
diff
changeset
|
511 |
finitizable_dataTs)) |
33192 | 512 |
end |
513 |
||
514 |
end; |