author | paulson <lp15@cam.ac.uk> |
Tue, 25 Jan 2022 14:13:33 +0000 | |
changeset 75008 | 43b3d5318d72 |
parent 66453 | cc19f7ca2ed6 |
permissions | -rw-r--r-- |
39186 | 1 |
theory IMP_2 |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
2 |
imports "HOL-Library.Predicate_Compile_Quickcheck" |
39186 | 3 |
begin |
4 |
||
63167 | 5 |
subsection \<open>IMP\<close> |
39186 | 6 |
|
63167 | 7 |
text \<open> |
39186 | 8 |
In this example, the state is one boolean variable and the commands are Skip, Ass, Seq, IF and While. |
63167 | 9 |
\<close> |
39186 | 10 |
|
42031 | 11 |
type_synonym var = unit |
12 |
type_synonym state = bool |
|
39186 | 13 |
|
58310 | 14 |
datatype com = |
39186 | 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'" |
|
45451
74515e8e6046
renaming example invocations: tester predicate_compile is renamed to smart_exhaustive
bulwahn
parents:
43686
diff
changeset
|
32 |
quickcheck[tester = smart_exhaustive, size = 2, iterations = 100, timeout = 600.0, expect = counterexample] |
39186 | 33 |
oops |
34 |
||
35 |
||
36 |
end |