| author | haftmann | 
| Thu, 08 Sep 2011 11:31:23 +0200 | |
| changeset 44839 | d19c677eb812 | 
| parent 43686 | bc7d63c7fd6f | 
| child 45451 | 74515e8e6046 | 
| permissions | -rw-r--r-- | 
| 39186 | 1 | theory IMP_1 | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40924diff
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, and IF. | |
| 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 | ||
| 20 | inductive exec :: "com => state => state => bool" where | |
| 21 | "exec Skip s s" | | |
| 22 | "exec (Ass e) s e" | | |
| 23 | "exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" | | |
| 24 | "s ==> exec c1 s t ==> exec (IF c1 c2) s t" | | |
| 25 | "\<not> s ==> exec c2 s t ==> exec (IF c1 c2) s t" | |
| 26 | ||
| 27 | lemma | |
| 28 | "exec c s s' ==> exec (Seq c c) s s'" | |
| 43686 
bc7d63c7fd6f
tuning options to avoid spurious isabelle test failures
 bulwahn parents: 
42397diff
changeset | 29 | quickcheck[tester = predicate_compile_wo_ff, size = 2, iterations = 100, timeout = 600.0, expect = counterexample] | 
| 39186 | 30 | oops | 
| 31 | ||
| 32 | ||
| 33 | end |