author | himmelma |
Thu, 28 May 2009 13:56:50 +0200 | |
changeset 31278 | 60a53b5af39c |
parent 31267 | 4a85a4afc97d |
child 31483 | 88210717bfc8 |
permissions | -rw-r--r-- |
29132 | 1 |
(* Author: Florian Haftmann, TU Muenchen *) |
26265 | 2 |
|
3 |
header {* A simple counterexample generator *} |
|
4 |
||
5 |
theory Quickcheck |
|
31203
5c8fb4fd67e0
moved Code_Index, Random and Quickcheck before Main
haftmann
parents:
31194
diff
changeset
|
6 |
imports Random Code_Eval |
31260 | 7 |
uses ("Tools/quickcheck_generators.ML") |
26265 | 8 |
begin |
9 |
||
31179 | 10 |
notation fcomp (infixl "o>" 60) |
11 |
notation scomp (infixl "o\<rightarrow>" 60) |
|
12 |
||
13 |
||
26265 | 14 |
subsection {* The @{text random} class *} |
15 |
||
28335 | 16 |
class random = typerep + |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
17 |
fixes random :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed" |
26265 | 18 |
|
26267
ba710daf77a7
added combinator for interpretation of construction of datatype
haftmann
parents:
26265
diff
changeset
|
19 |
|
31260 | 20 |
subsection {* Fundamental and numeric types*} |
31179 | 21 |
|
22 |
instantiation bool :: random |
|
23 |
begin |
|
24 |
||
25 |
definition |
|
31194 | 26 |
"random i = Random.range i o\<rightarrow> |
27 |
(\<lambda>k. Pair (if (k div 2 = 0) then Code_Eval.valtermify True else Code_Eval.valtermify False))" |
|
31179 | 28 |
|
29 |
instance .. |
|
30 |
||
31 |
end |
|
32 |
||
33 |
instantiation itself :: (typerep) random |
|
34 |
begin |
|
35 |
||
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
36 |
definition random_itself :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a itself \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where |
31194 | 37 |
"random_itself _ = Pair (Code_Eval.valtermify TYPE('a))" |
31179 | 38 |
|
39 |
instance .. |
|
40 |
||
41 |
end |
|
42 |
||
43 |
instantiation nat :: random |
|
44 |
begin |
|
45 |
||
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
46 |
definition random_nat :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (nat \<times> (unit \<Rightarrow> Code_Eval.term)) \<times> Random.seed" where |
31194 | 47 |
"random_nat i = Random.range (i + 1) o\<rightarrow> (\<lambda>k. Pair ( |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
48 |
let n = Code_Numeral.nat_of k |
31194 | 49 |
in (n, \<lambda>_. Code_Eval.term_of n)))" |
50 |
||
51 |
instance .. |
|
52 |
||
53 |
end |
|
54 |
||
55 |
instantiation int :: random |
|
56 |
begin |
|
57 |
||
58 |
definition |
|
59 |
"random i = Random.range (2 * i + 1) o\<rightarrow> (\<lambda>k. Pair ( |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31203
diff
changeset
|
60 |
let j = (if k \<ge> i then Code_Numeral.int_of (k - i) else - Code_Numeral.int_of (i - k)) |
31194 | 61 |
in (j, \<lambda>_. Code_Eval.term_of j)))" |
31179 | 62 |
|
63 |
instance .. |
|
64 |
||
30945 | 65 |
end |
31179 | 66 |
|
31260 | 67 |
|
68 |
subsection {* Complex generators *} |
|
69 |
||
70 |
definition collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where |
|
71 |
"collapse f = (f o\<rightarrow> id)" |
|
31223
87bde6b5f793
re-added corrected version of type copy quickcheck generator
haftmann
parents:
31211
diff
changeset
|
72 |
|
31260 | 73 |
definition beyond :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where |
74 |
"beyond k l = (if l > k then l else 0)" |
|
75 |
||
31267 | 76 |
lemma beyond_zero: |
77 |
"beyond k 0 = 0" |
|
78 |
by (simp add: beyond_def) |
|
79 |
||
31260 | 80 |
use "Tools/quickcheck_generators.ML" |
81 |
setup {* Quickcheck_Generators.setup *} |
|
82 |
||
83 |
code_reserved Quickcheck Quickcheck_Generators |
|
84 |
||
85 |
text {* Type @{typ "'a \<Rightarrow> 'b"} *} |
|
31223
87bde6b5f793
re-added corrected version of type copy quickcheck generator
haftmann
parents:
31211
diff
changeset
|
86 |
|
31260 | 87 |
axiomatization random_fun_aux :: "typerep \<Rightarrow> typerep \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> term) |
88 |
\<Rightarrow> (Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> (Random.seed \<Rightarrow> Random.seed \<times> Random.seed) |
|
89 |
\<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" |
|
31223
87bde6b5f793
re-added corrected version of type copy quickcheck generator
haftmann
parents:
31211
diff
changeset
|
90 |
|
31260 | 91 |
code_const random_fun_aux (Quickcheck "Quickcheck'_Generators.random'_fun") |
92 |
-- {* With enough criminal energy this can be abused to derive @{prop False}; |
|
93 |
for this reason we use a distinguished target @{text Quickcheck} |
|
94 |
not spoiling the regular trusted code generation *} |
|
31223
87bde6b5f793
re-added corrected version of type copy quickcheck generator
haftmann
parents:
31211
diff
changeset
|
95 |
|
31260 | 96 |
instantiation "fun" :: ("{eq, term_of}", "{type, random}") random |
97 |
begin |
|
31223
87bde6b5f793
re-added corrected version of type copy quickcheck generator
haftmann
parents:
31211
diff
changeset
|
98 |
|
31260 | 99 |
definition random_fun :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where |
100 |
"random n = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of (random n) Random.split_seed" |
|
101 |
||
102 |
instance .. |
|
31223
87bde6b5f793
re-added corrected version of type copy quickcheck generator
haftmann
parents:
31211
diff
changeset
|
103 |
|
87bde6b5f793
re-added corrected version of type copy quickcheck generator
haftmann
parents:
31211
diff
changeset
|
104 |
end |
31245 | 105 |
|
106 |
||
31267 | 107 |
hide (open) const collapse beyond |
108 |
||
31179 | 109 |
no_notation fcomp (infixl "o>" 60) |
110 |
no_notation scomp (infixl "o\<rightarrow>" 60) |
|
111 |
||
112 |
end |