| author | bulwahn | 
| Wed, 06 Jul 2011 13:52:42 +0200 | |
| changeset 43686 | bc7d63c7fd6f | 
| parent 42397 | 13798dcbdca5 | 
| child 45451 | 74515e8e6046 | 
| permissions | -rw-r--r-- | 
| 39186 | 1  | 
theory IMP_2  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40924 
diff
changeset
 | 
2  | 
imports "~~/src/HOL/Library/Predicate_Compile_Quickcheck"  | 
| 39186 | 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, IF and While.  | 
|
9  | 
*}  | 
|
10  | 
||
| 42031 | 11  | 
type_synonym var = unit  | 
12  | 
type_synonym state = bool  | 
|
| 39186 | 13  | 
|
14  | 
datatype com =  | 
|
15  | 
Skip |  | 
|
16  | 
Ass bool |  | 
|
17  | 
Seq com com |  | 
|
18  | 
IF com com |  | 
|
19  | 
While 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  | 
"\<not> s ==> exec (While c) s s" |  | 
|
28  | 
"s ==> exec c s s' ==> exec (While c) s' s'' ==> exec (While c) s s''"  | 
|
29  | 
||
30  | 
lemma  | 
|
31  | 
"exec c s s' ==> exec (Seq c c) s s'"  | 
|
| 
43686
 
bc7d63c7fd6f
tuning options to avoid spurious isabelle test failures
 
bulwahn 
parents: 
42397 
diff
changeset
 | 
32  | 
quickcheck[tester = predicate_compile_wo_ff, size = 2, iterations = 100, timeout = 600.0, expect = counterexample]  | 
| 39186 | 33  | 
oops  | 
34  | 
||
35  | 
||
36  | 
end  |