author | bulwahn |
Mon, 20 Sep 2010 09:26:19 +0200 | |
changeset 39544 | 3a07bbc264b2 |
parent 39249 | 9c866b248cb1 |
child 40924 | a9be7f26b4e6 |
permissions | -rw-r--r-- |
39186 | 1 |
theory IMP_1 |
2 |
imports "Predicate_Compile_Quickcheck" |
|
3 |
begin |
|
4 |
||
5 |
subsection {* IMP *} |
|
6 |
||
7 |
text {* |
|
8 |
In this example, the state is one boolean variable and the commands are Skip, Ass, Seq, and IF. |
|
9 |
*} |
|
10 |
||
11 |
types |
|
12 |
var = unit |
|
13 |
state = bool |
|
14 |
||
15 |
datatype com = |
|
16 |
Skip | |
|
17 |
Ass bool | |
|
18 |
Seq com com | |
|
19 |
IF com com |
|
20 |
||
21 |
inductive exec :: "com => state => state => bool" where |
|
22 |
"exec Skip s s" | |
|
23 |
"exec (Ass e) s e" | |
|
24 |
"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" | |
|
25 |
"s ==> exec c1 s t ==> exec (IF c1 c2) s t" | |
|
26 |
"\<not> s ==> exec c2 s t ==> exec (IF c1 c2) s t" |
|
27 |
||
28 |
lemma |
|
29 |
"exec c s s' ==> exec (Seq c c) s s'" |
|
39249
9c866b248cb1
increasing the number of iterations to ensure to find a counterexample by random generation
bulwahn
parents:
39186
diff
changeset
|
30 |
quickcheck[generator = predicate_compile_wo_ff, size = 2, iterations = 10, expect = counterexample] |
39186 | 31 |
oops |
32 |
||
33 |
||
34 |
end |