author | bulwahn |
Wed, 23 Sep 2009 16:20:13 +0200 | |
changeset 32672 | 90f3ce5d27ae |
parent 32670 | cc0bae788b7e |
child 32673 | d5db9cf85401 |
permissions | -rw-r--r-- |
31129
d2cead76fca2
split Predicate_Compile examples into separate theory
haftmann
parents:
31123
diff
changeset
|
1 |
theory Predicate_Compile_ex |
32318
bca7fd849829
improved use of context with cases rule in predicate compiler; predicate compiler based on Main for faster debugging
bulwahn
parents:
32317
diff
changeset
|
2 |
imports Main Predicate_Compile |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
3 |
begin |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
4 |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
5 |
inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
6 |
"even 0" |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
7 |
| "even n \<Longrightarrow> odd (Suc n)" |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
8 |
| "odd n \<Longrightarrow> even (Suc n)" |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
9 |
|
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31514
diff
changeset
|
10 |
code_pred even . |
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31514
diff
changeset
|
11 |
|
31514 | 12 |
thm odd.equation |
31123 | 13 |
thm even.equation |
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
14 |
|
31217 | 15 |
values "{x. even 2}" |
16 |
values "{x. odd 2}" |
|
17 |
values 10 "{n. even n}" |
|
18 |
values 10 "{n. odd n}" |
|
31195 | 19 |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
20 |
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where |
32340 | 21 |
"append [] xs xs" |
22 |
| "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)" |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
23 |
|
32340 | 24 |
code_pred append . |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
25 |
code_pred (inductify_all) (rpred) append . |
31123 | 26 |
|
27 |
thm append.equation |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
28 |
thm append.rpred_equation |
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
29 |
|
31217 | 30 |
values "{(ys, xs). append xs ys [0, Suc 0, 2]}" |
31 |
values "{zs. append [0, Suc 0, 2] [17, 8] zs}" |
|
32 |
values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0,5]}" |
|
31195 | 33 |
|
32340 | 34 |
inductive rev where |
35 |
"rev [] []" |
|
36 |
| "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys" |
|
37 |
||
38 |
code_pred rev . |
|
39 |
||
40 |
thm rev.equation |
|
41 |
||
42 |
values "{xs. rev [0, 1, 2, 3::nat] xs}" |
|
30972
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
43 |
|
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
44 |
inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
45 |
for f where |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
46 |
"partition f [] [] []" |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
47 |
| "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs" |
5b65835ccc92
some experiements towards user interface for predicate compiler
haftmann
parents:
30953
diff
changeset
|
48 |
| "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)" |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
49 |
|
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31514
diff
changeset
|
50 |
code_pred partition . |
31123 | 51 |
|
52 |
thm partition.equation |
|
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
53 |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
54 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
55 |
inductive member |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
56 |
for xs |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
57 |
where "x \<in> set xs ==> member xs x" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
58 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
59 |
lemma [code_pred_intros]: |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
60 |
"member (x#xs') x" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
61 |
by (auto intro: member.intros) |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
62 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
63 |
lemma [code_pred_intros]: |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
64 |
"member xs x ==> member (x'#xs) x" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
65 |
by (auto intro: member.intros elim!: member.cases) |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
66 |
(* strange bug must be repaired! *) |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
67 |
(* |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
68 |
code_pred member sorry |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
69 |
*) |
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
70 |
inductive is_even :: "nat \<Rightarrow> bool" |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
71 |
where |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
72 |
"n mod 2 = 0 \<Longrightarrow> is_even n" |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
73 |
|
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
74 |
code_pred is_even . |
31123 | 75 |
|
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
76 |
values 10 "{(ys, zs). partition is_even |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
77 |
[0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}" |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
78 |
values 10 "{zs. partition is_even zs [0, 2] [3, 5]}" |
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
79 |
values 10 "{zs. partition is_even zs [0, 7] [3, 5]}" |
31217 | 80 |
|
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31514
diff
changeset
|
81 |
lemma [code_pred_intros]: |
32340 | 82 |
"r a b \<Longrightarrow> tranclp r a b" |
83 |
"r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c" |
|
84 |
by auto |
|
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
85 |
|
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
86 |
code_pred tranclp |
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
87 |
proof - |
31580 | 88 |
case tranclp |
89 |
from this converse_tranclpE[OF this(1)] show thesis by metis |
|
31573
0047df9eb347
improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents:
31551
diff
changeset
|
90 |
qed |
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
91 |
(* |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
92 |
code_pred (inductify_all) (rpred) tranclp . |
31123 | 93 |
thm tranclp.equation |
32316
1d83ac469459
added generator compilation of higher-order predicates; refined mode analysis for generators; some tuning
bulwahn
parents:
32314
diff
changeset
|
94 |
thm tranclp.rpred_equation |
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
95 |
*) |
31217 | 96 |
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where |
97 |
"succ 0 1" |
|
98 |
| "succ m n \<Longrightarrow> succ (Suc m) (Suc n)" |
|
99 |
||
31550
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31514
diff
changeset
|
100 |
code_pred succ . |
b63d253ed9e2
code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents:
31514
diff
changeset
|
101 |
|
31217 | 102 |
thm succ.equation |
32340 | 103 |
|
104 |
values 10 "{(m, n). succ n m}" |
|
105 |
values "{m. succ 0 m}" |
|
106 |
values "{m. succ m 0}" |
|
107 |
||
32314
66bbad0bfef9
changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents:
32310
diff
changeset
|
108 |
(* FIXME: why does this not terminate? -- value chooses mode [] --> [1] and then starts enumerating all successors *) |
32355 | 109 |
|
31514 | 110 |
(* |
31217 | 111 |
values 20 "{n. tranclp succ 10 n}" |
112 |
values "{n. tranclp succ n 10}" |
|
113 |
values 20 "{(n, m). tranclp succ n m}" |
|
31514 | 114 |
*) |
31217 | 115 |
|
32424 | 116 |
subsection{* IMP *} |
117 |
||
118 |
types |
|
119 |
var = nat |
|
120 |
state = "int list" |
|
121 |
||
122 |
datatype com = |
|
123 |
Skip | |
|
124 |
Ass var "state => int" | |
|
125 |
Seq com com | |
|
126 |
IF "state => bool" com com | |
|
127 |
While "state => bool" com |
|
128 |
||
129 |
inductive exec :: "com => state => state => bool" where |
|
130 |
"exec Skip s s" | |
|
131 |
"exec (Ass x e) s (s[x := e(s)])" | |
|
132 |
"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" | |
|
133 |
"b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" | |
|
134 |
"~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" | |
|
135 |
"~b s ==> exec (While b c) s s" | |
|
136 |
"b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3" |
|
137 |
||
138 |
code_pred exec . |
|
139 |
||
140 |
values "{t. exec |
|
141 |
(While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1)))) |
|
142 |
[3,5] t}" |
|
143 |
||
144 |
||
32408 | 145 |
subsection{* CCS *} |
146 |
||
147 |
text{* This example formalizes finite CCS processes without communication or |
|
148 |
recursion. For simplicity, labels are natural numbers. *} |
|
149 |
||
150 |
datatype proc = nil | pre nat proc | or proc proc | par proc proc |
|
151 |
||
152 |
inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where |
|
153 |
"step (pre n p) n p" | |
|
154 |
"step p1 a q \<Longrightarrow> step (or p1 p2) a q" | |
|
155 |
"step p2 a q \<Longrightarrow> step (or p1 p2) a q" | |
|
156 |
"step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" | |
|
157 |
"step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)" |
|
158 |
||
159 |
code_pred step . |
|
160 |
||
161 |
inductive steps where |
|
162 |
"steps p [] p" | |
|
163 |
"step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r" |
|
164 |
||
165 |
code_pred steps . |
|
166 |
||
167 |
values 5 |
|
168 |
"{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}" |
|
169 |
||
170 |
(* FIXME |
|
171 |
values 3 "{(a,q). step (par nil nil) a q}" |
|
172 |
*) |
|
173 |
||
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
174 |
subsection {* divmod *} |
32579 | 175 |
|
176 |
inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where |
|
177 |
"k < l \<Longrightarrow> divmod_rel k l 0 k" |
|
178 |
| "k \<ge> l \<Longrightarrow> divmod_rel (k - l) l q r \<Longrightarrow> divmod_rel k l (Suc q) r" |
|
179 |
||
180 |
code_pred divmod_rel .. |
|
181 |
||
182 |
value [code] "Predicate.singleton (divmod_rel_1_2 1705 42)" |
|
183 |
||
32670 | 184 |
section {* Executing definitions *} |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
185 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
186 |
definition Min |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
187 |
where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
188 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
189 |
code_pred (inductify_all) Min . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
190 |
|
32670 | 191 |
subsection {* Examples with lists *} |
192 |
||
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
193 |
inductive filterP for Pa where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
194 |
"(filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) [] []" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
195 |
| "[| (res::'a list) = (y::'a) # (resa::'a list); (filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) (xt::'a list) resa; Pa y |] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
196 |
==> filterP Pa (y # xt) res" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
197 |
| "[| (filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) (xt::'a list) (res::'a list); ~ Pa (y::'a) |] ==> filterP Pa (y # xt) res" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
198 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
199 |
(* |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
200 |
code_pred (inductify_all) (rpred) filterP . |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
201 |
thm filterP.rpred_equation |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
202 |
*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
203 |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
204 |
code_pred (inductify_all) lexord . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
205 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
206 |
thm lexord.equation |
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
207 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
208 |
lemma "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
209 |
(*quickcheck[generator=pred_compile]*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
210 |
oops |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
211 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
212 |
lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
213 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
214 |
code_pred (inductify_all) lexn . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
215 |
thm lexn.equation |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
216 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
217 |
code_pred (inductify_all) lenlex . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
218 |
thm lenlex.equation |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
219 |
(* |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
220 |
code_pred (inductify_all) (rpred) lenlex . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
221 |
thm lenlex.rpred_equation |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
222 |
*) |
32670 | 223 |
thm lists.intros |
224 |
code_pred (inductify_all) lists . |
|
225 |
||
226 |
thm lists.equation |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
227 |
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
228 |
datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
229 |
fun height :: "'a tree => nat" where |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
230 |
"height ET = 0" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
231 |
| "height (MKT x l r h) = max (height l) (height r) + 1" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
232 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
233 |
consts avl :: "'a tree => bool" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
234 |
primrec |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
235 |
"avl ET = True" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
236 |
"avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and> |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
237 |
h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
238 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
239 |
code_pred (inductify_all) avl . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
240 |
thm avl.equation |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
241 |
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
242 |
lemma [code_pred_inline]: "bot_fun_inst.bot_fun == (\<lambda>(y::'a::type). False)" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
243 |
unfolding bot_fun_inst.bot_fun[symmetric] bot_bool_eq[symmetric] bot_fun_eq by simp |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
244 |
|
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
245 |
fun set_of |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
246 |
where |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
247 |
"set_of ET = {}" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
248 |
| "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
249 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
250 |
fun is_ord |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
251 |
where |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
252 |
"is_ord ET = True" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
253 |
| "is_ord (MKT n l r h) = |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
254 |
((\<forall>n' \<in> set_of l. n' < n) \<and> (\<forall>n' \<in> set_of r. n < n') \<and> is_ord l \<and> is_ord r)" |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
255 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
256 |
declare Un_def[code_pred_def] |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
257 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
258 |
code_pred (inductify_all) set_of . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
259 |
thm set_of.equation |
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
260 |
(* FIXME *) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
261 |
(* |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
262 |
code_pred (inductify_all) is_ord . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
263 |
thm is_ord.equation |
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
264 |
*) |
32668
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
265 |
section {* Definitions about Relations *} |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
266 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
267 |
code_pred (inductify_all) converse . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
268 |
thm converse.equation |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
269 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
270 |
code_pred (inductify_all) Domain . |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
271 |
thm Domain.equation |
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
272 |
|
b2de45007537
added first prototype of the extended predicate compiler
bulwahn
parents:
32665
diff
changeset
|
273 |
|
32669
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
274 |
section {* Context Free Grammar *} |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
275 |
|
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
276 |
datatype alphabet = a | b |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
277 |
|
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
278 |
inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
279 |
"[] \<in> S\<^isub>1" |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
280 |
| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1" |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
281 |
| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1" |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
282 |
| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1" |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
283 |
| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1" |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
284 |
| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1" |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
285 |
|
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
286 |
code_pred (inductify_all) S\<^isub>1p . |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
287 |
|
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
288 |
thm S\<^isub>1p.equation |
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
289 |
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
290 |
theorem S\<^isub>1_sound: |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
291 |
"w \<in> S\<^isub>1 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
292 |
quickcheck[generator=pred_compile] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
293 |
oops |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
294 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
295 |
inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
296 |
"[] \<in> S\<^isub>2" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
297 |
| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
298 |
| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
299 |
| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
300 |
| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
301 |
| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
302 |
(* |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
303 |
code_pred (inductify_all) (rpred) S\<^isub>2 . |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
304 |
ML {* Predicate_Compile_Core.intros_of @{theory} @{const_name "B\<^isub>2"} *} |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
305 |
*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
306 |
theorem S\<^isub>2_sound: |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
307 |
"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
308 |
(*quickcheck[generator=SML]*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
309 |
quickcheck[generator=pred_compile, size=15, iterations=100] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
310 |
oops |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
311 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
312 |
inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
313 |
"[] \<in> S\<^isub>3" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
314 |
| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
315 |
| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
316 |
| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
317 |
| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
318 |
| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
319 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
320 |
ML {* reset Predicate_Compile_Core.do_proofs *} |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
321 |
(* |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
322 |
code_pred (inductify_all) S\<^isub>3 . |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
323 |
*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
324 |
theorem S\<^isub>3_sound: |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
325 |
"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
326 |
quickcheck[generator=pred_compile, size=10, iterations=1] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
327 |
oops |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
328 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
329 |
lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
330 |
quickcheck[size=10, generator = pred_compile] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
331 |
oops |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
332 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
333 |
theorem S\<^isub>3_complete: |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
334 |
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
335 |
(*quickcheck[generator=SML]*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
336 |
quickcheck[generator=pred_compile, size=10, iterations=100] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
337 |
oops |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
338 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
339 |
inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
340 |
"[] \<in> S\<^isub>4" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
341 |
| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
342 |
| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
343 |
| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
344 |
| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
345 |
| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
346 |
| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
347 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
348 |
theorem S\<^isub>4_sound: |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
349 |
"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
350 |
quickcheck[generator = pred_compile, size=2, iterations=1] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
351 |
sorry |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
352 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
353 |
theorem S\<^isub>4_complete: |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
354 |
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
355 |
quickcheck[generator = pred_compile, size=5, iterations=1] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
356 |
sorry |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
357 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
358 |
theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete: |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
359 |
"w \<in> S\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
360 |
"w \<in> A\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
361 |
"w \<in> B\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
362 |
(*quickcheck[generator = pred_compile, size=5, iterations=1]*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
363 |
sorry |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
364 |
|
32669
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
365 |
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
366 |
section {* Lambda *} |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
367 |
datatype type = |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
368 |
Atom nat |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
369 |
| Fun type type (infixr "\<Rightarrow>" 200) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
370 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
371 |
datatype dB = |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
372 |
Var nat |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
373 |
| App dB dB (infixl "\<degree>" 200) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
374 |
| Abs type dB |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
375 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
376 |
primrec |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
377 |
nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
378 |
where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
379 |
"[]\<langle>i\<rangle> = None" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
380 |
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
381 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
382 |
(* |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
383 |
inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
384 |
where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
385 |
"nth_el' (x # xs) 0 x" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
386 |
| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
387 |
*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
388 |
inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
389 |
where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
390 |
Var [intro!]: "nth_el env x = Some T \<Longrightarrow> env \<turnstile> Var x : T" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
391 |
| Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
392 |
(* | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" *) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
393 |
| App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
394 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
395 |
primrec |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
396 |
lift :: "[dB, nat] => dB" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
397 |
where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
398 |
"lift (Var i) k = (if i < k then Var i else Var (i + 1))" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
399 |
| "lift (s \<degree> t) k = lift s k \<degree> lift t k" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
400 |
| "lift (Abs T s) k = Abs T (lift s (k + 1))" |
32669
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
401 |
|
32672
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
402 |
primrec |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
403 |
subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
404 |
where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
405 |
subst_Var: "(Var i)[s/k] = |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
406 |
(if k < i then Var (i - 1) else if i = k then s else Var i)" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
407 |
| subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
408 |
| subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
409 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
410 |
inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
411 |
where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
412 |
beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
413 |
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
414 |
| appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
415 |
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
416 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
417 |
lemma "Gamma \<turnstile> t : T \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> Gamma \<turnstile> t' : T" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
418 |
quickcheck[generator = pred_compile, size = 10, iterations = 1000] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
419 |
oops |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
420 |
(* FIXME *) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
421 |
(* |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
422 |
inductive test for P where |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
423 |
"[| filter P vs = res |] |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
424 |
==> test P vs res" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
425 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
426 |
code_pred test . |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
427 |
*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
428 |
(* |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
429 |
export_code test_for_1_yields_1_2 in SML file - |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
430 |
code_pred (inductify_all) (rpred) test . |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
431 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
432 |
thm test.equation |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
433 |
*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
434 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
435 |
ML {*set Toplevel.debug *} |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
436 |
|
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
437 |
lemma filter_eq_ConsD: |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
438 |
"filter P ys = x#xs \<Longrightarrow> |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
439 |
\<exists>us vs. ys = ts @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs" |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
440 |
(*quickcheck[generator = pred_compile]*) |
90f3ce5d27ae
added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents:
32670
diff
changeset
|
441 |
oops |
32669
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
442 |
|
462b1dd67a58
added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents:
32668
diff
changeset
|
443 |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
444 |
end |