src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy
author bulwahn
Fri, 30 Oct 2009 09:55:15 +0100
changeset 33375 fd3e861f8d31
parent 33257 95186fb5653c
child 33405 5c1928d5db38
permissions -rw-r--r--
renamed rpred to random
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
     1
theory Predicate_Compile_Quickcheck_ex
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
     2
imports Predicate_Compile_Quickcheck
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
     3
  Predicate_Compile_Alternative_Defs
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
     4
begin
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
     5
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
     6
section {* Sets *}
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
     7
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
     8
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
     9
section {* Context Free Grammar *}
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    10
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    11
datatype alphabet = a | b
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    12
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    13
inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    14
  "[] \<in> S\<^isub>1"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    15
| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    16
| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    17
| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    18
| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    19
| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    20
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    21
theorem S\<^isub>1_sound:
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    22
"w \<in> S\<^isub>1p \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    23
(*quickcheck[generator=predicate_compile, size=15]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    24
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    25
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    26
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    27
inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    28
  "[] \<in> S\<^isub>2"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    29
| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    30
| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    31
| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    32
| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    33
| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    34
33375
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33257
diff changeset
    35
code_pred [inductify, random] S\<^isub>2 .
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33257
diff changeset
    36
thm S\<^isub>2.random_equation
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33257
diff changeset
    37
thm A\<^isub>2.random_equation
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33257
diff changeset
    38
thm B\<^isub>2.random_equation
33257
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    39
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    40
values [random] 10 "{x. S\<^isub>2 x}"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    41
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    42
lemma "w \<in> S\<^isub>2 ==> w \<noteq> [] ==> w \<noteq> [b, a] ==> w \<in> {}"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    43
quickcheck[generator=predicate_compile]
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    44
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    45
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    46
lemma "[x <- w. x = a] = []"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    47
quickcheck[generator=predicate_compile]
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    48
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    49
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    50
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    51
lemma "length ([x \<leftarrow> w. x = a]) = (0::nat)"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    52
(*quickcheck[generator=predicate_compile]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    53
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    54
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    55
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    56
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    57
lemma
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    58
"w \<in> S\<^isub>2 ==> length [x \<leftarrow> w. x = a] < Suc (Suc 0)"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    59
(*quickcheck[generator=predicate_compile]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    60
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    61
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    62
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    63
theorem S\<^isub>2_sound:
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    64
"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    65
(*quickcheck[generator=predicate_compile, size=15, iterations=100]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    66
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    67
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    68
inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    69
  "[] \<in> S\<^isub>3"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    70
| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    71
| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    72
| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    73
| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    74
| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    75
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    76
code_pred [inductify] S\<^isub>3 .
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    77
thm S\<^isub>3.equation
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    78
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    79
values 10 "{x. S\<^isub>3 x}"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    80
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    81
lemma S\<^isub>3_sound:
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    82
"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    83
(*quickcheck[generator=predicate_compile, size=10, iterations=1]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    84
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    85
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    86
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    87
lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    88
(*quickcheck[size=10, generator = pred_compile]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    89
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    90
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    91
theorem S\<^isub>3_complete:
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    92
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    93
(*quickcheck[generator=SML]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    94
(*quickcheck[generator=predicate_compile, size=10, iterations=100]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    95
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    96
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    97
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    98
inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
    99
  "[] \<in> S\<^isub>4"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   100
| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   101
| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   102
| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   103
| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   104
| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   105
| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   106
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   107
theorem S\<^isub>4_sound:
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   108
"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   109
(*quickcheck[generator = predicate_compile, size=2, iterations=1]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   110
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   111
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   112
theorem S\<^isub>4_complete:
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   113
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   114
(*quickcheck[generator = pred_compile, size=5, iterations=1]*)
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   115
oops
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   116
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   117
95186fb5653c added examples for quickcheck prototype
bulwahn
parents:
diff changeset
   118
end