author | haftmann |
Mon, 31 May 2021 20:27:45 +0000 | |
changeset 73793 | 26c0ccf17f31 |
parent 72607 | feebdaa346e5 |
child 80932 | 261cd8722677 |
permissions | -rw-r--r-- |
62979 | 1 |
(* Title: HOL/Quickcheck_Exhaustive.thy |
2 |
Author: Lukas Bulwahn, TU Muenchen |
|
3 |
*) |
|
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
4 |
|
60758 | 5 |
section \<open>A simple counterexample generator performing exhaustive testing\<close> |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
6 |
|
41918 | 7 |
theory Quickcheck_Exhaustive |
51126
df86080de4cb
reform of predicate compiler / quickcheck theories:
haftmann
parents:
49948
diff
changeset
|
8 |
imports Quickcheck_Random |
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
46417
diff
changeset
|
9 |
keywords "quickcheck_generator" :: thy_decl |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
10 |
begin |
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
11 |
|
62979 | 12 |
subsection \<open>Basic operations for exhaustive generators\<close> |
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
13 |
|
62979 | 14 |
definition orelse :: "'a option \<Rightarrow> 'a option \<Rightarrow> 'a option" (infixr "orelse" 55) |
15 |
where [code_unfold]: "x orelse y = (case x of Some x' \<Rightarrow> Some x' | None \<Rightarrow> y)" |
|
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
16 |
|
62979 | 17 |
|
18 |
subsection \<open>Exhaustive generator type classes\<close> |
|
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
19 |
|
41916
80060d5f864a
renaming constants in Quickcheck_Exhaustive theory
bulwahn
parents:
41915
diff
changeset
|
20 |
class exhaustive = term_of + |
62979 | 21 |
fixes exhaustive :: "('a \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option" |
22 |
||
42310
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
23 |
class full_exhaustive = term_of + |
62979 | 24 |
fixes full_exhaustive :: |
25 |
"('a \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option" |
|
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
26 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
27 |
instantiation natural :: full_exhaustive |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
28 |
begin |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
29 |
|
62979 | 30 |
function full_exhaustive_natural' :: |
31 |
"(natural \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> |
|
32 |
natural \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option" |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
33 |
where "full_exhaustive_natural' f d i = |
42304 | 34 |
(if d < i then None |
62979 | 35 |
else (f (i, \<lambda>_. Code_Evaluation.term_of i)) orelse (full_exhaustive_natural' f d (i + 1)))" |
42304 | 36 |
by pat_completeness auto |
37 |
||
38 |
termination |
|
62979 | 39 |
by (relation "measure (\<lambda>(_, d, i). nat_of_natural (d + 1 - i))") (auto simp add: less_natural_def) |
42304 | 40 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
41 |
definition "full_exhaustive f d = full_exhaustive_natural' f d 0" |
42304 | 42 |
|
42310
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
43 |
instance .. |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
44 |
|
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
45 |
end |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
46 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
47 |
instantiation natural :: exhaustive |
42310
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
48 |
begin |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
49 |
|
62979 | 50 |
function exhaustive_natural' :: |
51 |
"(natural \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option" |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
52 |
where "exhaustive_natural' f d i = |
41916
80060d5f864a
renaming constants in Quickcheck_Exhaustive theory
bulwahn
parents:
41915
diff
changeset
|
53 |
(if d < i then None |
62979 | 54 |
else (f i orelse exhaustive_natural' f d (i + 1)))" |
41231
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
55 |
by pat_completeness auto |
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
56 |
|
42304 | 57 |
termination |
62979 | 58 |
by (relation "measure (\<lambda>(_, d, i). nat_of_natural (d + 1 - i))") (auto simp add: less_natural_def) |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
59 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
60 |
definition "exhaustive f d = exhaustive_natural' f d 0" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
61 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
62 |
instance .. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
63 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
64 |
end |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
65 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
66 |
instantiation integer :: exhaustive |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
67 |
begin |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
68 |
|
62979 | 69 |
function exhaustive_integer' :: |
70 |
"(integer \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> integer \<Rightarrow> integer \<Rightarrow> (bool \<times> term list) option" |
|
71 |
where "exhaustive_integer' f d i = |
|
72 |
(if d < i then None else (f i orelse exhaustive_integer' f d (i + 1)))" |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
73 |
by pat_completeness auto |
41231
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
74 |
|
62979 | 75 |
termination |
76 |
by (relation "measure (\<lambda>(_, d, i). nat_of_integer (d + 1 - i))") |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
77 |
(auto simp add: less_integer_def nat_of_integer_def) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
78 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
79 |
definition "exhaustive f d = exhaustive_integer' f (integer_of_natural d) (- (integer_of_natural d))" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
80 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
81 |
instance .. |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
82 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
83 |
end |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
84 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
85 |
instantiation integer :: full_exhaustive |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
86 |
begin |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
87 |
|
62979 | 88 |
function full_exhaustive_integer' :: |
89 |
"(integer \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> |
|
90 |
integer \<Rightarrow> integer \<Rightarrow> (bool \<times> term list) option" |
|
91 |
where "full_exhaustive_integer' f d i = |
|
92 |
(if d < i then None |
|
93 |
else |
|
94 |
(case f (i, \<lambda>_. Code_Evaluation.term_of i) of |
|
95 |
Some t \<Rightarrow> Some t |
|
96 |
| None \<Rightarrow> full_exhaustive_integer' f d (i + 1)))" |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
97 |
by pat_completeness auto |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
98 |
|
62979 | 99 |
termination |
100 |
by (relation "measure (\<lambda>(_, d, i). nat_of_integer (d + 1 - i))") |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
101 |
(auto simp add: less_integer_def nat_of_integer_def) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
102 |
|
62979 | 103 |
definition "full_exhaustive f d = |
104 |
full_exhaustive_integer' f (integer_of_natural d) (- (integer_of_natural d))" |
|
41231
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
105 |
|
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
106 |
instance .. |
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
107 |
|
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
108 |
end |
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
109 |
|
41916
80060d5f864a
renaming constants in Quickcheck_Exhaustive theory
bulwahn
parents:
41915
diff
changeset
|
110 |
instantiation nat :: exhaustive |
41231
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
111 |
begin |
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
112 |
|
62979 | 113 |
definition "exhaustive f d = exhaustive (\<lambda>x. f (nat_of_natural x)) d" |
42304 | 114 |
|
42310
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
115 |
instance .. |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
116 |
|
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
117 |
end |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
118 |
|
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
119 |
instantiation nat :: full_exhaustive |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
120 |
begin |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
121 |
|
62979 | 122 |
definition "full_exhaustive f d = |
123 |
full_exhaustive (\<lambda>(x, xt). f (nat_of_natural x, \<lambda>_. Code_Evaluation.term_of (nat_of_natural x))) d" |
|
41231
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
124 |
|
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
125 |
instance .. |
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
126 |
|
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
127 |
end |
2e901158675e
adding exhaustive tester instances for numeric types: code_numeral, nat, rat and real
bulwahn
parents:
41178
diff
changeset
|
128 |
|
41916
80060d5f864a
renaming constants in Quickcheck_Exhaustive theory
bulwahn
parents:
41915
diff
changeset
|
129 |
instantiation int :: exhaustive |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
130 |
begin |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
131 |
|
62979 | 132 |
function exhaustive_int' :: |
133 |
"(int \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> (bool \<times> term list) option" |
|
134 |
where "exhaustive_int' f d i = |
|
135 |
(if d < i then None else (f i orelse exhaustive_int' f d (i + 1)))" |
|
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
136 |
by pat_completeness auto |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
137 |
|
62979 | 138 |
termination |
139 |
by (relation "measure (\<lambda>(_, d, i). nat (d + 1 - i))") auto |
|
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
140 |
|
62979 | 141 |
definition "exhaustive f d = |
142 |
exhaustive_int' f (int_of_integer (integer_of_natural d)) |
|
143 |
(- (int_of_integer (integer_of_natural d)))" |
|
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
144 |
|
42310
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
145 |
instance .. |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
146 |
|
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
147 |
end |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
148 |
|
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
149 |
instantiation int :: full_exhaustive |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
150 |
begin |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
151 |
|
62979 | 152 |
function full_exhaustive_int' :: |
153 |
"(int \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> |
|
154 |
int \<Rightarrow> int \<Rightarrow> (bool \<times> term list) option" |
|
155 |
where "full_exhaustive_int' f d i = |
|
156 |
(if d < i then None |
|
157 |
else |
|
158 |
(case f (i, \<lambda>_. Code_Evaluation.term_of i) of |
|
159 |
Some t \<Rightarrow> Some t |
|
160 |
| None \<Rightarrow> full_exhaustive_int' f d (i + 1)))" |
|
42304 | 161 |
by pat_completeness auto |
162 |
||
62979 | 163 |
termination |
164 |
by (relation "measure (\<lambda>(_, d, i). nat (d + 1 - i))") auto |
|
42304 | 165 |
|
62979 | 166 |
definition "full_exhaustive f d = |
167 |
full_exhaustive_int' f (int_of_integer (integer_of_natural d)) |
|
168 |
(- (int_of_integer (integer_of_natural d)))" |
|
42304 | 169 |
|
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
170 |
instance .. |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
171 |
|
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
172 |
end |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
173 |
|
41916
80060d5f864a
renaming constants in Quickcheck_Exhaustive theory
bulwahn
parents:
41915
diff
changeset
|
174 |
instantiation prod :: (exhaustive, exhaustive) exhaustive |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
175 |
begin |
40899
ef6fde932f4c
improving readability of Smallcheck theory; adding constant orelse to improve performance of the function package
bulwahn
parents:
40639
diff
changeset
|
176 |
|
62979 | 177 |
definition "exhaustive f d = exhaustive (\<lambda>x. exhaustive (\<lambda>y. f ((x, y))) d) d" |
42304 | 178 |
|
42310
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
179 |
instance .. |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
180 |
|
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
181 |
end |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
182 |
|
72607 | 183 |
context |
184 |
includes term_syntax |
|
185 |
begin |
|
186 |
||
187 |
definition |
|
62979 | 188 |
[code_unfold]: "valtermify_pair x y = |
189 |
Code_Evaluation.valtermify (Pair :: 'a::typerep \<Rightarrow> 'b::typerep \<Rightarrow> 'a \<times> 'b) {\<cdot>} x {\<cdot>} y" |
|
46307
ec8f975c059b
shortened definitions by adding some termify constants
bulwahn
parents:
46305
diff
changeset
|
190 |
|
72607 | 191 |
end |
192 |
||
42310
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
193 |
instantiation prod :: (full_exhaustive, full_exhaustive) full_exhaustive |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
194 |
begin |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
195 |
|
62979 | 196 |
definition "full_exhaustive f d = |
197 |
full_exhaustive (\<lambda>x. full_exhaustive (\<lambda>y. f (valtermify_pair x y)) d) d" |
|
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
198 |
|
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
199 |
instance .. |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
200 |
|
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
201 |
end |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
202 |
|
46193
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
203 |
instantiation set :: (exhaustive) exhaustive |
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
204 |
begin |
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
205 |
|
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
206 |
fun exhaustive_set |
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
207 |
where |
62979 | 208 |
"exhaustive_set f i = |
209 |
(if i = 0 then None |
|
210 |
else |
|
211 |
f {} orelse |
|
212 |
exhaustive_set |
|
213 |
(\<lambda>A. f A orelse exhaustive (\<lambda>x. if x \<in> A then None else f (insert x A)) (i - 1)) (i - 1))" |
|
46193
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
214 |
|
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
215 |
instance .. |
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
216 |
|
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
217 |
end |
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
218 |
|
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
219 |
instantiation set :: (full_exhaustive) full_exhaustive |
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
220 |
begin |
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
221 |
|
62979 | 222 |
fun full_exhaustive_set |
46193
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
223 |
where |
62979 | 224 |
"full_exhaustive_set f i = |
225 |
(if i = 0 then None |
|
226 |
else |
|
227 |
f valterm_emptyset orelse |
|
228 |
full_exhaustive_set |
|
229 |
(\<lambda>A. f A orelse Quickcheck_Exhaustive.full_exhaustive |
|
230 |
(\<lambda>x. if fst x \<in> fst A then None else f (valtermify_insert x A)) (i - 1)) (i - 1))" |
|
46193
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
231 |
|
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
232 |
instance .. |
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
233 |
|
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
234 |
end |
55a4769d0abe
adding exhaustive instances for type constructor set
bulwahn
parents:
45925
diff
changeset
|
235 |
|
62979 | 236 |
instantiation "fun" :: ("{equal,exhaustive}", exhaustive) exhaustive |
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
237 |
begin |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
238 |
|
62979 | 239 |
fun exhaustive_fun' :: |
240 |
"(('a \<Rightarrow> 'b) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> natural \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option" |
|
42304 | 241 |
where |
62979 | 242 |
"exhaustive_fun' f i d = |
243 |
(exhaustive (\<lambda>b. f (\<lambda>_. b)) d) orelse |
|
244 |
(if i > 1 then |
|
245 |
exhaustive_fun' |
|
246 |
(\<lambda>g. exhaustive (\<lambda>a. exhaustive (\<lambda>b. f (g(a := b))) d) d) (i - 1) d else None)" |
|
42304 | 247 |
|
62979 | 248 |
definition exhaustive_fun :: |
249 |
"(('a \<Rightarrow> 'b) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option" |
|
250 |
where "exhaustive_fun f d = exhaustive_fun' f d d" |
|
42304 | 251 |
|
42310
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
252 |
instance .. |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
253 |
|
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
254 |
end |
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
255 |
|
62979 | 256 |
definition [code_unfold]: |
257 |
"valtermify_absdummy = |
|
258 |
(\<lambda>(v, t). |
|
259 |
(\<lambda>_::'a. v, |
|
260 |
\<lambda>u::unit. Code_Evaluation.Abs (STR ''x'') (Typerep.typerep TYPE('a::typerep)) (t ())))" |
|
46307
ec8f975c059b
shortened definitions by adding some termify constants
bulwahn
parents:
46305
diff
changeset
|
261 |
|
72607 | 262 |
context |
263 |
includes term_syntax |
|
264 |
begin |
|
265 |
||
266 |
definition |
|
62979 | 267 |
[code_unfold]: "valtermify_fun_upd g a b = |
268 |
Code_Evaluation.valtermify |
|
269 |
(fun_upd :: ('a::typerep \<Rightarrow> 'b::typerep) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b) {\<cdot>} g {\<cdot>} a {\<cdot>} b" |
|
46307
ec8f975c059b
shortened definitions by adding some termify constants
bulwahn
parents:
46305
diff
changeset
|
270 |
|
72607 | 271 |
end |
272 |
||
62979 | 273 |
instantiation "fun" :: ("{equal,full_exhaustive}", full_exhaustive) full_exhaustive |
42310
c664cc5cc5e9
splitting exhaustive and full_exhaustive into separate type classes
bulwahn
parents:
42305
diff
changeset
|
274 |
begin |
42304 | 275 |
|
62979 | 276 |
fun full_exhaustive_fun' :: |
277 |
"(('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> |
|
278 |
natural \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option" |
|
42304 | 279 |
where |
62979 | 280 |
"full_exhaustive_fun' f i d = |
281 |
full_exhaustive (\<lambda>v. f (valtermify_absdummy v)) d orelse |
|
282 |
(if i > 1 then |
|
283 |
full_exhaustive_fun' |
|
284 |
(\<lambda>g. full_exhaustive |
|
285 |
(\<lambda>a. full_exhaustive (\<lambda>b. f (valtermify_fun_upd g a b)) d) d) (i - 1) d |
|
286 |
else None)" |
|
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
287 |
|
62979 | 288 |
definition full_exhaustive_fun :: |
289 |
"(('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> |
|
290 |
natural \<Rightarrow> (bool \<times> term list) option" |
|
291 |
where "full_exhaustive_fun f d = full_exhaustive_fun' f d d" |
|
40639
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
292 |
|
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
293 |
instance .. |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
294 |
|
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
295 |
end |
f1f0e6adca0a
adding function generation to SmallCheck; activating exhaustive search space testing
bulwahn
parents:
40620
diff
changeset
|
296 |
|
60758 | 297 |
subsubsection \<open>A smarter enumeration scheme for functions over finite datatypes\<close> |
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
298 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
299 |
class check_all = enum + term_of + |
62979 | 300 |
fixes check_all :: "('a \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> (bool * term list) option" |
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
301 |
fixes enum_term_of :: "'a itself \<Rightarrow> unit \<Rightarrow> term list" |
62979 | 302 |
|
303 |
fun check_all_n_lists :: "('a::check_all list \<times> (unit \<Rightarrow> term list) \<Rightarrow> |
|
304 |
(bool \<times> term list) option) \<Rightarrow> natural \<Rightarrow> (bool * term list) option" |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
305 |
where |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
306 |
"check_all_n_lists f n = |
62979 | 307 |
(if n = 0 then f ([], (\<lambda>_. [])) |
308 |
else check_all (\<lambda>(x, xt). |
|
309 |
check_all_n_lists (\<lambda>(xs, xst). f ((x # xs), (\<lambda>_. (xt () # xst ())))) (n - 1)))" |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
310 |
|
72607 | 311 |
context |
312 |
includes term_syntax |
|
313 |
begin |
|
314 |
||
315 |
definition |
|
62979 | 316 |
[code_unfold]: "termify_fun_upd g a b = |
317 |
(Code_Evaluation.termify |
|
318 |
(fun_upd :: ('a::typerep \<Rightarrow> 'b::typerep) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b) <\<cdot>> g <\<cdot>> a <\<cdot>> b)" |
|
46307
ec8f975c059b
shortened definitions by adding some termify constants
bulwahn
parents:
46305
diff
changeset
|
319 |
|
72607 | 320 |
end |
321 |
||
62979 | 322 |
definition mk_map_term :: |
323 |
"(unit \<Rightarrow> typerep) \<Rightarrow> (unit \<Rightarrow> typerep) \<Rightarrow> |
|
324 |
(unit \<Rightarrow> term list) \<Rightarrow> (unit \<Rightarrow> term list) \<Rightarrow> unit \<Rightarrow> term" |
|
325 |
where "mk_map_term T1 T2 domm rng = |
|
326 |
(\<lambda>_. |
|
327 |
let |
|
328 |
T1 = T1 (); |
|
329 |
T2 = T2 (); |
|
330 |
update_term = |
|
331 |
(\<lambda>g (a, b). |
|
332 |
Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.App |
|
333 |
(Code_Evaluation.Const (STR ''Fun.fun_upd'') |
|
334 |
(Typerep.Typerep (STR ''fun'') [Typerep.Typerep (STR ''fun'') [T1, T2], |
|
335 |
Typerep.Typerep (STR ''fun'') [T1, |
|
336 |
Typerep.Typerep (STR ''fun'') [T2, Typerep.Typerep (STR ''fun'') [T1, T2]]]])) |
|
337 |
g) a) b) |
|
338 |
in |
|
339 |
List.foldl update_term |
|
340 |
(Code_Evaluation.Abs (STR ''x'') T1 |
|
341 |
(Code_Evaluation.Const (STR ''HOL.undefined'') T2)) (zip (domm ()) (rng ())))" |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
342 |
|
62979 | 343 |
instantiation "fun" :: ("{equal,check_all}", check_all) check_all |
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
344 |
begin |
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
345 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
346 |
definition |
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
347 |
"check_all f = |
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
348 |
(let |
62979 | 349 |
mk_term = |
350 |
mk_map_term |
|
351 |
(\<lambda>_. Typerep.typerep (TYPE('a))) |
|
352 |
(\<lambda>_. Typerep.typerep (TYPE('b))) |
|
353 |
(enum_term_of (TYPE('a))); |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
354 |
enum = (Enum.enum :: 'a list) |
62979 | 355 |
in |
356 |
check_all_n_lists |
|
67091 | 357 |
(\<lambda>(ys, yst). f (the \<circ> map_of (zip enum ys), mk_term yst)) |
62979 | 358 |
(natural_of_nat (length enum)))" |
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
359 |
|
62979 | 360 |
definition enum_term_of_fun :: "('a \<Rightarrow> 'b) itself \<Rightarrow> unit \<Rightarrow> term list" |
361 |
where "enum_term_of_fun = |
|
362 |
(\<lambda>_ _. |
|
363 |
let |
|
364 |
enum_term_of_a = enum_term_of (TYPE('a)); |
|
365 |
mk_term = |
|
366 |
mk_map_term |
|
367 |
(\<lambda>_. Typerep.typerep (TYPE('a))) |
|
368 |
(\<lambda>_. Typerep.typerep (TYPE('b))) |
|
369 |
enum_term_of_a |
|
370 |
in |
|
371 |
map (\<lambda>ys. mk_term (\<lambda>_. ys) ()) |
|
372 |
(List.n_lists (length (enum_term_of_a ())) (enum_term_of (TYPE('b)) ())))" |
|
373 |
||
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
374 |
instance .. |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
375 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
376 |
end |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
377 |
|
72607 | 378 |
context |
379 |
includes term_syntax |
|
380 |
begin |
|
381 |
||
382 |
fun check_all_subsets :: |
|
62979 | 383 |
"(('a::typerep) set \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> |
384 |
('a \<times> (unit \<Rightarrow> term)) list \<Rightarrow> (bool \<times> term list) option" |
|
46305 | 385 |
where |
386 |
"check_all_subsets f [] = f valterm_emptyset" |
|
62979 | 387 |
| "check_all_subsets f (x # xs) = |
388 |
check_all_subsets (\<lambda>s. case f s of Some ts \<Rightarrow> Some ts | None \<Rightarrow> f (valtermify_insert x s)) xs" |
|
46305 | 389 |
|
72607 | 390 |
definition |
62979 | 391 |
[code_unfold]: "term_emptyset = Code_Evaluation.termify ({} :: ('a::typerep) set)" |
46305 | 392 |
|
72607 | 393 |
definition |
62979 | 394 |
[code_unfold]: "termify_insert x s = |
395 |
Code_Evaluation.termify (insert :: ('a::typerep) \<Rightarrow> 'a set \<Rightarrow> 'a set) <\<cdot>> x <\<cdot>> s" |
|
396 |
||
72607 | 397 |
definition setify :: "('a::typerep) itself \<Rightarrow> term list \<Rightarrow> term" |
46305 | 398 |
where |
62979 | 399 |
"setify T ts = foldr (termify_insert T) ts (term_emptyset T)" |
46305 | 400 |
|
72607 | 401 |
end |
402 |
||
46305 | 403 |
instantiation set :: (check_all) check_all |
404 |
begin |
|
405 |
||
406 |
definition |
|
407 |
"check_all_set f = |
|
62979 | 408 |
check_all_subsets f |
409 |
(zip (Enum.enum :: 'a list) |
|
410 |
(map (\<lambda>a. \<lambda>u :: unit. a) (Quickcheck_Exhaustive.enum_term_of (TYPE ('a)) ())))" |
|
46305 | 411 |
|
62979 | 412 |
definition enum_term_of_set :: "'a set itself \<Rightarrow> unit \<Rightarrow> term list" |
413 |
where "enum_term_of_set _ _ = |
|
65956
639eb3617a86
reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents:
64670
diff
changeset
|
414 |
map (setify (TYPE('a))) (subseqs (Quickcheck_Exhaustive.enum_term_of (TYPE('a)) ()))" |
46305 | 415 |
|
416 |
instance .. |
|
417 |
||
418 |
end |
|
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
419 |
|
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
420 |
instantiation unit :: check_all |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
421 |
begin |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
422 |
|
62979 | 423 |
definition "check_all f = f (Code_Evaluation.valtermify ())" |
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
424 |
|
62979 | 425 |
definition enum_term_of_unit :: "unit itself \<Rightarrow> unit \<Rightarrow> term list" |
426 |
where "enum_term_of_unit = (\<lambda>_ _. [Code_Evaluation.term_of ()])" |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
427 |
|
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
428 |
instance .. |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
429 |
|
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
430 |
end |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
431 |
|
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
432 |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
433 |
instantiation bool :: check_all |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
434 |
begin |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
435 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
436 |
definition |
62979 | 437 |
"check_all f = |
438 |
(case f (Code_Evaluation.valtermify False) of |
|
439 |
Some x' \<Rightarrow> Some x' |
|
440 |
| None \<Rightarrow> f (Code_Evaluation.valtermify True))" |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
441 |
|
62979 | 442 |
definition enum_term_of_bool :: "bool itself \<Rightarrow> unit \<Rightarrow> term list" |
443 |
where "enum_term_of_bool = (\<lambda>_ _. map Code_Evaluation.term_of (Enum.enum :: bool list))" |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
444 |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
445 |
instance .. |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
446 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
447 |
end |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
448 |
|
72607 | 449 |
context |
450 |
includes term_syntax |
|
451 |
begin |
|
452 |
||
453 |
definition [code_unfold]: |
|
62979 | 454 |
"termify_pair x y = |
455 |
Code_Evaluation.termify (Pair :: 'a::typerep \<Rightarrow> 'b :: typerep \<Rightarrow> 'a * 'b) <\<cdot>> x <\<cdot>> y" |
|
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
456 |
|
72607 | 457 |
end |
458 |
||
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
459 |
instantiation prod :: (check_all, check_all) check_all |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
460 |
begin |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
461 |
|
62979 | 462 |
definition "check_all f = check_all (\<lambda>x. check_all (\<lambda>y. f (valtermify_pair x y)))" |
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
463 |
|
62979 | 464 |
definition enum_term_of_prod :: "('a * 'b) itself \<Rightarrow> unit \<Rightarrow> term list" |
465 |
where "enum_term_of_prod = |
|
466 |
(\<lambda>_ _. |
|
467 |
map (\<lambda>(x, y). termify_pair TYPE('a) TYPE('b) x y) |
|
468 |
(List.product (enum_term_of (TYPE('a)) ()) (enum_term_of (TYPE('b)) ())))" |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
469 |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
470 |
instance .. |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
471 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
472 |
end |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
473 |
|
72607 | 474 |
context |
475 |
includes term_syntax |
|
476 |
begin |
|
477 |
||
478 |
definition |
|
62979 | 479 |
[code_unfold]: "valtermify_Inl x = |
480 |
Code_Evaluation.valtermify (Inl :: 'a::typerep \<Rightarrow> 'a + 'b :: typerep) {\<cdot>} x" |
|
481 |
||
72607 | 482 |
definition |
62979 | 483 |
[code_unfold]: "valtermify_Inr x = |
484 |
Code_Evaluation.valtermify (Inr :: 'b::typerep \<Rightarrow> 'a::typerep + 'b) {\<cdot>} x" |
|
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
485 |
|
72607 | 486 |
end |
487 |
||
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
488 |
instantiation sum :: (check_all, check_all) check_all |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
489 |
begin |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
490 |
|
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
491 |
definition |
62979 | 492 |
"check_all f = check_all (\<lambda>a. f (valtermify_Inl a)) orelse check_all (\<lambda>b. f (valtermify_Inr b))" |
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
493 |
|
62979 | 494 |
definition enum_term_of_sum :: "('a + 'b) itself \<Rightarrow> unit \<Rightarrow> term list" |
495 |
where "enum_term_of_sum = |
|
496 |
(\<lambda>_ _. |
|
497 |
let |
|
498 |
T1 = Typerep.typerep (TYPE('a)); |
|
499 |
T2 = Typerep.typerep (TYPE('b)) |
|
500 |
in |
|
501 |
map |
|
502 |
(Code_Evaluation.App (Code_Evaluation.Const (STR ''Sum_Type.Inl'') |
|
503 |
(Typerep.Typerep (STR ''fun'') [T1, Typerep.Typerep (STR ''Sum_Type.sum'') [T1, T2]]))) |
|
504 |
(enum_term_of (TYPE('a)) ()) @ |
|
505 |
map |
|
506 |
(Code_Evaluation.App (Code_Evaluation.Const (STR ''Sum_Type.Inr'') |
|
507 |
(Typerep.Typerep (STR ''fun'') [T2, Typerep.Typerep (STR ''Sum_Type.sum'') [T1, T2]]))) |
|
508 |
(enum_term_of (TYPE('b)) ()))" |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
509 |
|
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
510 |
instance .. |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
511 |
|
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
512 |
end |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
513 |
|
64670
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
514 |
instantiation char :: check_all |
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
515 |
begin |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
516 |
|
64670
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
517 |
primrec check_all_char' :: |
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
518 |
"(char \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> char list \<Rightarrow> (bool \<times> term list) option" |
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
519 |
where "check_all_char' f [] = None" |
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
520 |
| "check_all_char' f (c # cs) = f (c, \<lambda>_. Code_Evaluation.term_of c) |
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
521 |
orelse check_all_char' f cs" |
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
522 |
|
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
523 |
definition check_all_char :: |
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
524 |
"(char \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> (bool \<times> term list) option" |
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
525 |
where "check_all f = check_all_char' f Enum.enum" |
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
526 |
|
62979 | 527 |
definition enum_term_of_char :: "char itself \<Rightarrow> unit \<Rightarrow> term list" |
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
528 |
where |
62979 | 529 |
"enum_term_of_char = (\<lambda>_ _. map Code_Evaluation.term_of (Enum.enum :: char list))" |
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
530 |
|
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
531 |
instance .. |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
532 |
|
64670
f77b946d18aa
restored instance for char, which got ancidentally lost in b3f2b8c906a6
haftmann
parents:
62979
diff
changeset
|
533 |
end |
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
534 |
|
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
535 |
instantiation option :: (check_all) check_all |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
536 |
begin |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
537 |
|
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
538 |
definition |
62979 | 539 |
"check_all f = |
540 |
f (Code_Evaluation.valtermify (None :: 'a option)) orelse |
|
541 |
check_all |
|
542 |
(\<lambda>(x, t). |
|
543 |
f |
|
544 |
(Some x, |
|
545 |
\<lambda>_. Code_Evaluation.App |
|
546 |
(Code_Evaluation.Const (STR ''Option.option.Some'') |
|
547 |
(Typerep.Typerep (STR ''fun'') |
|
548 |
[Typerep.typerep TYPE('a), |
|
549 |
Typerep.Typerep (STR ''Option.option'') [Typerep.typerep TYPE('a)]])) (t ())))" |
|
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
550 |
|
62979 | 551 |
definition enum_term_of_option :: "'a option itself \<Rightarrow> unit \<Rightarrow> term list" |
552 |
where "enum_term_of_option = |
|
553 |
(\<lambda> _ _. |
|
554 |
Code_Evaluation.term_of (None :: 'a option) # |
|
555 |
(map |
|
556 |
(Code_Evaluation.App |
|
557 |
(Code_Evaluation.Const (STR ''Option.option.Some'') |
|
558 |
(Typerep.Typerep (STR ''fun'') |
|
559 |
[Typerep.typerep TYPE('a), |
|
560 |
Typerep.Typerep (STR ''Option.option'') [Typerep.typerep TYPE('a)]]))) |
|
561 |
(enum_term_of (TYPE('a)) ())))" |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
562 |
|
41105
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
563 |
instance .. |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
564 |
|
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
565 |
end |
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
566 |
|
a76ee71c3313
adding check_all instances for a few more finite types in smallcheck
bulwahn
parents:
41104
diff
changeset
|
567 |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
568 |
instantiation Enum.finite_1 :: check_all |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
569 |
begin |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
570 |
|
62979 | 571 |
definition "check_all f = f (Code_Evaluation.valtermify Enum.finite_1.a\<^sub>1)" |
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
572 |
|
62979 | 573 |
definition enum_term_of_finite_1 :: "Enum.finite_1 itself \<Rightarrow> unit \<Rightarrow> term list" |
574 |
where "enum_term_of_finite_1 = (\<lambda>_ _. [Code_Evaluation.term_of Enum.finite_1.a\<^sub>1])" |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
575 |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
576 |
instance .. |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
577 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
578 |
end |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
579 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
580 |
instantiation Enum.finite_2 :: check_all |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
581 |
begin |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
582 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
583 |
definition |
62979 | 584 |
"check_all f = |
585 |
(f (Code_Evaluation.valtermify Enum.finite_2.a\<^sub>1) orelse |
|
586 |
f (Code_Evaluation.valtermify Enum.finite_2.a\<^sub>2))" |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
587 |
|
62979 | 588 |
definition enum_term_of_finite_2 :: "Enum.finite_2 itself \<Rightarrow> unit \<Rightarrow> term list" |
589 |
where "enum_term_of_finite_2 = |
|
590 |
(\<lambda>_ _. map Code_Evaluation.term_of (Enum.enum :: Enum.finite_2 list))" |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
591 |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
592 |
instance .. |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
593 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
594 |
end |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
595 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
596 |
instantiation Enum.finite_3 :: check_all |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
597 |
begin |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
598 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
599 |
definition |
62979 | 600 |
"check_all f = |
601 |
(f (Code_Evaluation.valtermify Enum.finite_3.a\<^sub>1) orelse |
|
602 |
f (Code_Evaluation.valtermify Enum.finite_3.a\<^sub>2) orelse |
|
603 |
f (Code_Evaluation.valtermify Enum.finite_3.a\<^sub>3))" |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
604 |
|
62979 | 605 |
definition enum_term_of_finite_3 :: "Enum.finite_3 itself \<Rightarrow> unit \<Rightarrow> term list" |
606 |
where "enum_term_of_finite_3 = |
|
607 |
(\<lambda>_ _. map Code_Evaluation.term_of (Enum.enum :: Enum.finite_3 list))" |
|
41177
810a885decee
added enum_term_of to correct present nested functions
bulwahn
parents:
41105
diff
changeset
|
608 |
|
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
609 |
instance .. |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
610 |
|
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
611 |
end |
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
612 |
|
46417
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
613 |
instantiation Enum.finite_4 :: check_all |
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
614 |
begin |
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
615 |
|
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
616 |
definition |
62979 | 617 |
"check_all f = |
618 |
f (Code_Evaluation.valtermify Enum.finite_4.a\<^sub>1) orelse |
|
619 |
f (Code_Evaluation.valtermify Enum.finite_4.a\<^sub>2) orelse |
|
620 |
f (Code_Evaluation.valtermify Enum.finite_4.a\<^sub>3) orelse |
|
621 |
f (Code_Evaluation.valtermify Enum.finite_4.a\<^sub>4)" |
|
46417
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
622 |
|
62979 | 623 |
definition enum_term_of_finite_4 :: "Enum.finite_4 itself \<Rightarrow> unit \<Rightarrow> term list" |
624 |
where "enum_term_of_finite_4 = |
|
625 |
(\<lambda>_ _. map Code_Evaluation.term_of (Enum.enum :: Enum.finite_4 list))" |
|
46417
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
626 |
|
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
627 |
instance .. |
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
628 |
|
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
629 |
end |
1a68fcb80b62
beautifying definitions of check_all and adding instance for finite_4
bulwahn
parents:
46311
diff
changeset
|
630 |
|
60758 | 631 |
subsection \<open>Bounded universal quantifiers\<close> |
41085
a549ff1d4070
adding a smarter enumeration scheme for finite functions
bulwahn
parents:
40915
diff
changeset
|
632 |
|
42195
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42117
diff
changeset
|
633 |
class bounded_forall = |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
634 |
fixes bounded_forall :: "('a \<Rightarrow> bool) \<Rightarrow> natural \<Rightarrow> bool" |
42195
1e7b62c93f5d
adding an exhaustive validator for quickcheck's batch validating; moving strip_imp; minimal setup for bounded_forall
bulwahn
parents:
42117
diff
changeset
|
635 |
|
62979 | 636 |
|
60758 | 637 |
subsection \<open>Fast exhaustive combinators\<close> |
42305
494c31fdec95
theory definitions for fast exhaustive quickcheck compilation
bulwahn
parents:
42304
diff
changeset
|
638 |
|
494c31fdec95
theory definitions for fast exhaustive quickcheck compilation
bulwahn
parents:
42304
diff
changeset
|
639 |
class fast_exhaustive = term_of + |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
640 |
fixes fast_exhaustive :: "('a \<Rightarrow> unit) \<Rightarrow> natural \<Rightarrow> unit" |
42305
494c31fdec95
theory definitions for fast exhaustive quickcheck compilation
bulwahn
parents:
42304
diff
changeset
|
641 |
|
62979 | 642 |
axiomatization throw_Counterexample :: "term list \<Rightarrow> unit" |
643 |
axiomatization catch_Counterexample :: "unit \<Rightarrow> term list option" |
|
42305
494c31fdec95
theory definitions for fast exhaustive quickcheck compilation
bulwahn
parents:
42304
diff
changeset
|
644 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
645 |
code_printing |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
646 |
constant throw_Counterexample \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
647 |
(Quickcheck) "raise (Exhaustive'_Generators.Counterexample _)" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
51143
diff
changeset
|
648 |
| constant catch_Counterexample \<rightharpoonup> |
62979 | 649 |
(Quickcheck) "(((_); NONE) handle Exhaustive'_Generators.Counterexample ts \<Rightarrow> SOME ts)" |
650 |
||
42305
494c31fdec95
theory definitions for fast exhaustive quickcheck compilation
bulwahn
parents:
42304
diff
changeset
|
651 |
|
60758 | 652 |
subsection \<open>Continuation passing style functions as plus monad\<close> |
62979 | 653 |
|
654 |
type_synonym 'a cps = "('a \<Rightarrow> term list option) \<Rightarrow> term list option" |
|
45450 | 655 |
|
656 |
definition cps_empty :: "'a cps" |
|
62979 | 657 |
where "cps_empty = (\<lambda>cont. None)" |
45450 | 658 |
|
62979 | 659 |
definition cps_single :: "'a \<Rightarrow> 'a cps" |
660 |
where "cps_single v = (\<lambda>cont. cont v)" |
|
45450 | 661 |
|
62979 | 662 |
definition cps_bind :: "'a cps \<Rightarrow> ('a \<Rightarrow> 'b cps) \<Rightarrow> 'b cps" |
663 |
where "cps_bind m f = (\<lambda>cont. m (\<lambda>a. (f a) cont))" |
|
45450 | 664 |
|
62979 | 665 |
definition cps_plus :: "'a cps \<Rightarrow> 'a cps \<Rightarrow> 'a cps" |
666 |
where "cps_plus a b = (\<lambda>c. case a c of None \<Rightarrow> b c | Some x \<Rightarrow> Some x)" |
|
667 |
||
668 |
definition cps_if :: "bool \<Rightarrow> unit cps" |
|
669 |
where "cps_if b = (if b then cps_single () else cps_empty)" |
|
45450 | 670 |
|
62979 | 671 |
definition cps_not :: "unit cps \<Rightarrow> unit cps" |
672 |
where "cps_not n = (\<lambda>c. case n (\<lambda>u. Some []) of None \<Rightarrow> c () | Some _ \<Rightarrow> None)" |
|
45450 | 673 |
|
62979 | 674 |
type_synonym 'a pos_bound_cps = |
675 |
"('a \<Rightarrow> (bool * term list) option) \<Rightarrow> natural \<Rightarrow> (bool * term list) option" |
|
45450 | 676 |
|
677 |
definition pos_bound_cps_empty :: "'a pos_bound_cps" |
|
62979 | 678 |
where "pos_bound_cps_empty = (\<lambda>cont i. None)" |
45450 | 679 |
|
62979 | 680 |
definition pos_bound_cps_single :: "'a \<Rightarrow> 'a pos_bound_cps" |
681 |
where "pos_bound_cps_single v = (\<lambda>cont i. cont v)" |
|
45450 | 682 |
|
62979 | 683 |
definition pos_bound_cps_bind :: "'a pos_bound_cps \<Rightarrow> ('a \<Rightarrow> 'b pos_bound_cps) \<Rightarrow> 'b pos_bound_cps" |
684 |
where "pos_bound_cps_bind m f = (\<lambda>cont i. if i = 0 then None else (m (\<lambda>a. (f a) cont i) (i - 1)))" |
|
45450 | 685 |
|
62979 | 686 |
definition pos_bound_cps_plus :: "'a pos_bound_cps \<Rightarrow> 'a pos_bound_cps \<Rightarrow> 'a pos_bound_cps" |
687 |
where "pos_bound_cps_plus a b = (\<lambda>c i. case a c i of None \<Rightarrow> b c i | Some x \<Rightarrow> Some x)" |
|
45450 | 688 |
|
62979 | 689 |
definition pos_bound_cps_if :: "bool \<Rightarrow> unit pos_bound_cps" |
690 |
where "pos_bound_cps_if b = (if b then pos_bound_cps_single () else pos_bound_cps_empty)" |
|
45450 | 691 |
|
58350
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
blanchet
parents:
58334
diff
changeset
|
692 |
datatype (plugins only: code extraction) (dead 'a) unknown = |
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
blanchet
parents:
58334
diff
changeset
|
693 |
Unknown | Known 'a |
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
blanchet
parents:
58334
diff
changeset
|
694 |
|
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
blanchet
parents:
58334
diff
changeset
|
695 |
datatype (plugins only: code extraction) (dead 'a) three_valued = |
919149921e46
added 'extraction' plugins -- this might help 'HOL-Proofs'
blanchet
parents:
58334
diff
changeset
|
696 |
Unknown_value | Value 'a | No_value |
45450 | 697 |
|
62979 | 698 |
type_synonym 'a neg_bound_cps = |
699 |
"('a unknown \<Rightarrow> term list three_valued) \<Rightarrow> natural \<Rightarrow> term list three_valued" |
|
45450 | 700 |
|
701 |
definition neg_bound_cps_empty :: "'a neg_bound_cps" |
|
62979 | 702 |
where "neg_bound_cps_empty = (\<lambda>cont i. No_value)" |
703 |
||
704 |
definition neg_bound_cps_single :: "'a \<Rightarrow> 'a neg_bound_cps" |
|
705 |
where "neg_bound_cps_single v = (\<lambda>cont i. cont (Known v))" |
|
45450 | 706 |
|
62979 | 707 |
definition neg_bound_cps_bind :: "'a neg_bound_cps \<Rightarrow> ('a \<Rightarrow> 'b neg_bound_cps) \<Rightarrow> 'b neg_bound_cps" |
708 |
where "neg_bound_cps_bind m f = |
|
709 |
(\<lambda>cont i. |
|
710 |
if i = 0 then cont Unknown |
|
711 |
else m (\<lambda>a. case a of Unknown \<Rightarrow> cont Unknown | Known a' \<Rightarrow> f a' cont i) (i - 1))" |
|
45450 | 712 |
|
62979 | 713 |
definition neg_bound_cps_plus :: "'a neg_bound_cps \<Rightarrow> 'a neg_bound_cps \<Rightarrow> 'a neg_bound_cps" |
714 |
where "neg_bound_cps_plus a b = |
|
715 |
(\<lambda>c i. |
|
716 |
case a c i of |
|
717 |
No_value \<Rightarrow> b c i |
|
718 |
| Value x \<Rightarrow> Value x |
|
719 |
| Unknown_value \<Rightarrow> |
|
720 |
(case b c i of |
|
721 |
No_value \<Rightarrow> Unknown_value |
|
722 |
| Value x \<Rightarrow> Value x |
|
723 |
| Unknown_value \<Rightarrow> Unknown_value))" |
|
45450 | 724 |
|
62979 | 725 |
definition neg_bound_cps_if :: "bool \<Rightarrow> unit neg_bound_cps" |
726 |
where "neg_bound_cps_if b = (if b then neg_bound_cps_single () else neg_bound_cps_empty)" |
|
45450 | 727 |
|
62979 | 728 |
definition neg_bound_cps_not :: "unit pos_bound_cps \<Rightarrow> unit neg_bound_cps" |
729 |
where "neg_bound_cps_not n = |
|
730 |
(\<lambda>c i. case n (\<lambda>u. Some (True, [])) i of None \<Rightarrow> c (Known ()) | Some _ \<Rightarrow> No_value)" |
|
731 |
||
732 |
definition pos_bound_cps_not :: "unit neg_bound_cps \<Rightarrow> unit pos_bound_cps" |
|
733 |
where "pos_bound_cps_not n = |
|
734 |
(\<lambda>c i. case n (\<lambda>u. Value []) i of No_value \<Rightarrow> c () | Value _ \<Rightarrow> None | Unknown_value \<Rightarrow> None)" |
|
735 |
||
45450 | 736 |
|
60758 | 737 |
subsection \<open>Defining generators for any first-order data type\<close> |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
738 |
|
45697
3b7c35a08723
more stable introduction of the internally used unknown term
bulwahn
parents:
45684
diff
changeset
|
739 |
axiomatization unknown :: 'a |
3b7c35a08723
more stable introduction of the internally used unknown term
bulwahn
parents:
45684
diff
changeset
|
740 |
|
3b7c35a08723
more stable introduction of the internally used unknown term
bulwahn
parents:
45684
diff
changeset
|
741 |
notation (output) unknown ("?") |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
742 |
|
69605 | 743 |
ML_file \<open>Tools/Quickcheck/exhaustive_generators.ML\<close> |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
744 |
|
43882
05d5696f177f
renaming quickcheck_tester to quickcheck_batch_tester; tuned
bulwahn
parents:
42310
diff
changeset
|
745 |
declare [[quickcheck_batch_tester = exhaustive]] |
40915
a4c956d1f91f
declaring quickcheck testers as default after their setup
bulwahn
parents:
40914
diff
changeset
|
746 |
|
62979 | 747 |
|
60758 | 748 |
subsection \<open>Defining generators for abstract types\<close> |
45925 | 749 |
|
69605 | 750 |
ML_file \<open>Tools/Quickcheck/abstract_generators.ML\<close> |
45925 | 751 |
|
47203 | 752 |
hide_fact (open) orelse_def |
62979 | 753 |
no_notation orelse (infixr "orelse" 55) |
45818
53a697f5454a
hiding constants and facts in the Quickcheck_Exhaustive and Quickcheck_Narrowing theory;
bulwahn
parents:
45750
diff
changeset
|
754 |
|
62979 | 755 |
hide_const valtermify_absdummy valtermify_fun_upd |
756 |
valterm_emptyset valtermify_insert |
|
757 |
valtermify_pair valtermify_Inl valtermify_Inr |
|
46307
ec8f975c059b
shortened definitions by adding some termify constants
bulwahn
parents:
46305
diff
changeset
|
758 |
termify_fun_upd term_emptyset termify_insert termify_pair setify |
46305 | 759 |
|
45818
53a697f5454a
hiding constants and facts in the Quickcheck_Exhaustive and Quickcheck_Narrowing theory;
bulwahn
parents:
45750
diff
changeset
|
760 |
hide_const (open) |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
761 |
exhaustive full_exhaustive |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
762 |
exhaustive_int' full_exhaustive_int' |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
763 |
exhaustive_integer' full_exhaustive_integer' |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
764 |
exhaustive_natural' full_exhaustive_natural' |
45818
53a697f5454a
hiding constants and facts in the Quickcheck_Exhaustive and Quickcheck_Narrowing theory;
bulwahn
parents:
45750
diff
changeset
|
765 |
throw_Counterexample catch_Counterexample |
53a697f5454a
hiding constants and facts in the Quickcheck_Exhaustive and Quickcheck_Narrowing theory;
bulwahn
parents:
45750
diff
changeset
|
766 |
check_all enum_term_of |
46305 | 767 |
orelse unknown mk_map_term check_all_n_lists check_all_subsets |
40420
552563ea3304
adding code and theory for smallvalue generators, but do not setup the interpretation yet
bulwahn
parents:
diff
changeset
|
768 |
|
45450 | 769 |
hide_type (open) cps pos_bound_cps neg_bound_cps unknown three_valued |
62979 | 770 |
|
45450 | 771 |
hide_const (open) cps_empty cps_single cps_bind cps_plus cps_if cps_not |
62979 | 772 |
pos_bound_cps_empty pos_bound_cps_single pos_bound_cps_bind |
773 |
pos_bound_cps_plus pos_bound_cps_if pos_bound_cps_not |
|
774 |
neg_bound_cps_empty neg_bound_cps_single neg_bound_cps_bind |
|
775 |
neg_bound_cps_plus neg_bound_cps_if neg_bound_cps_not |
|
45450 | 776 |
Unknown Known Unknown_value Value No_value |
777 |
||
778 |
end |