| author | wenzelm | 
| Tue, 26 Jan 2021 23:34:40 +0100 | |
| changeset 73194 | c0d6d57a9a31 | 
| parent 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 39186 | 1 | theory IMP_4 | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
63167diff
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 a list of integers and the commands are Skip, Ass, Seq, IF and While. | 
| 63167 | 9 | \<close> | 
| 39186 | 10 | |
| 42463 | 11 | type_synonym var = nat | 
| 12 | type_synonym state = "int list" | |
| 39186 | 13 | |
| 58310 | 14 | datatype com = | 
| 39186 | 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: 
43686diff
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: 
42463diff
changeset | 36 | end |