src/HOL/Predicate_Compile_Examples/IMP_4.thy
author haftmann
Wed, 08 Dec 2010 13:34:50 +0100
changeset 41075 4bed56dc95fb
parent 40924 a9be7f26b4e6
child 41413 64cd30d6b0b8
permissions -rw-r--r--
primitive definitions of bot/top/inf/sup for bool and fun are named with canonical suffix `_def` rather than `_eq`
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39186
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
     1
theory IMP_4
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
     2
imports Predicate_Compile_Quickcheck
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
     3
begin
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
     4
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
     5
subsection {* IMP *}
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
     6
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
     7
text {*
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
     8
  In this example, the state is a list of integers and the commands are Skip, Ass, Seq, IF and While.
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
     9
*}
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    10
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    11
types
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    12
  var = nat
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    13
  state = "int list"
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    14
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    15
datatype com =
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    16
  Skip |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    17
  Ass var "int" |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    18
  Seq com com |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    19
  IF "state list" com com |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    20
  While "state list" com
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    21
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    22
inductive exec :: "com => state => state => bool" where
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    23
  "exec Skip s s" |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    24
  "exec (Ass x e) s (s[x := e])" |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    25
  "exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    26
  "s \<in> set b ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    27
  "s \<notin> set b ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    28
  "s \<notin> set b ==> exec (While b c) s s" |
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    29
  "s1 \<in> set b ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    30
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    31
lemma
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    32
  "exec c s s' ==> exec (Seq c c) s s'"
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    33
  nitpick (* nitpick fails here! *)
40924
a9be7f26b4e6 adapting predicate_compile_quickcheck
bulwahn
parents: 39186
diff changeset
    34
  quickcheck[tester = predicate_compile_wo_ff, size=2, iterations=10, expect=counterexample]
39186
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    35
oops
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    36
475856793715 adding IMP quickcheck examples
bulwahn
parents:
diff changeset
    37
end