author | huffman |
Sun, 25 Mar 2012 20:15:39 +0200 | |
changeset 47108 | 2a1953f0d20d |
parent 45451 | 74515e8e6046 |
child 58249 | 180f1b3508ed |
permissions | -rw-r--r-- |
39186 | 1 |
theory IMP_4 |
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 a list of integers and the commands are Skip, Ass, Seq, IF and While. |
|
9 |
*} |
|
10 |
||
42463 | 11 |
type_synonym var = nat |
12 |
type_synonym state = "int list" |
|
39186 | 13 |
|
14 |
datatype com = |
|
15 |
Skip | |
|
16 |
Ass var "int" | |
|
17 |
Seq com com | |
|
18 |
IF "state list" com com | |
|
19 |
While "state list" com |
|
20 |
||
21 |
inductive exec :: "com => state => state => bool" where |
|
22 |
"exec Skip s s" | |
|
23 |
"exec (Ass x e) s (s[x := e])" | |
|
24 |
"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" | |
|
25 |
"s \<in> set b ==> exec c1 s t ==> exec (IF b c1 c2) s t" | |
|
26 |
"s \<notin> set b ==> exec c2 s t ==> exec (IF b c1 c2) s t" | |
|
27 |
"s \<notin> set b ==> exec (While b c) s s" | |
|
28 |
"s1 \<in> set b ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3" |
|
29 |
||
30 |
lemma |
|
31 |
"exec c s s' ==> exec (Seq c c) s s'" |
|
32 |
nitpick (* nitpick fails here! *) |
|
45451
74515e8e6046
renaming example invocations: tester predicate_compile is renamed to smart_exhaustive
bulwahn
parents:
43686
diff
changeset
|
33 |
quickcheck[tester = smart_exhaustive, size=2, iterations=100, expect=counterexample] |
39186 | 34 |
oops |
35 |
||
43686
bc7d63c7fd6f
tuning options to avoid spurious isabelle test failures
bulwahn
parents:
42463
diff
changeset
|
36 |
end |