author | wenzelm |
Wed, 29 Dec 2010 17:34:41 +0100 | |
changeset 41413 | 64cd30d6b0b8 |
parent 40924 | a9be7f26b4e6 |
child 41956 | c15ef1b85035 |
permissions | -rw-r--r-- |
39189 | 1 |
theory Reg_Exp_Example |
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
40924
diff
changeset
|
2 |
imports "~~/src/HOL/Library/Predicate_Compile_Quickcheck" Code_Prolog |
39188 | 3 |
begin |
4 |
||
5 |
text {* An example from the experiments from SmallCheck (http://www.cs.york.ac.uk/fp/smallcheck/) *} |
|
6 |
text {* The example (original in Haskell) was imported with Haskabelle and |
|
7 |
manually slightly adapted. |
|
8 |
*} |
|
9 |
||
10 |
datatype Nat = Zer |
|
11 |
| Suc Nat |
|
12 |
||
13 |
fun sub :: "Nat \<Rightarrow> Nat \<Rightarrow> Nat" |
|
14 |
where |
|
15 |
"sub x y = (case y of |
|
16 |
Zer \<Rightarrow> x |
|
17 |
| Suc y' \<Rightarrow> case x of |
|
18 |
Zer \<Rightarrow> Zer |
|
19 |
| Suc x' \<Rightarrow> sub x' y')" |
|
20 |
||
21 |
datatype Sym = N0 |
|
22 |
| N1 Sym |
|
23 |
||
24 |
datatype RE = Sym Sym |
|
25 |
| Or RE RE |
|
26 |
| Seq RE RE |
|
27 |
| And RE RE |
|
28 |
| Star RE |
|
29 |
| Empty |
|
30 |
||
31 |
||
32 |
function accepts :: "RE \<Rightarrow> Sym list \<Rightarrow> bool" and |
|
33 |
seqSplit :: "RE \<Rightarrow> RE \<Rightarrow> Sym list \<Rightarrow> Sym list \<Rightarrow> bool" and |
|
34 |
seqSplit'' :: "RE \<Rightarrow> RE \<Rightarrow> Sym list \<Rightarrow> Sym list \<Rightarrow> bool" and |
|
35 |
seqSplit' :: "RE \<Rightarrow> RE \<Rightarrow> Sym list \<Rightarrow> Sym list \<Rightarrow> bool" |
|
36 |
where |
|
37 |
"accepts re ss = (case re of |
|
38 |
Sym n \<Rightarrow> case ss of |
|
39 |
Nil \<Rightarrow> False |
|
40 |
| (n' # ss') \<Rightarrow> n = n' & List.null ss' |
|
41 |
| Or re1 re2 \<Rightarrow> accepts re1 ss | accepts re2 ss |
|
42 |
| Seq re1 re2 \<Rightarrow> seqSplit re1 re2 Nil ss |
|
43 |
| And re1 re2 \<Rightarrow> accepts re1 ss & accepts re2 ss |
|
44 |
| Star re' \<Rightarrow> case ss of |
|
45 |
Nil \<Rightarrow> True |
|
46 |
| (s # ss') \<Rightarrow> seqSplit re' re [s] ss' |
|
47 |
| Empty \<Rightarrow> List.null ss)" |
|
48 |
| "seqSplit re1 re2 ss2 ss = (seqSplit'' re1 re2 ss2 ss | seqSplit' re1 re2 ss2 ss)" |
|
49 |
| "seqSplit'' re1 re2 ss2 ss = (accepts re1 ss2 & accepts re2 ss)" |
|
50 |
| "seqSplit' re1 re2 ss2 ss = (case ss of |
|
51 |
Nil \<Rightarrow> False |
|
52 |
| (n # ss') \<Rightarrow> seqSplit re1 re2 (ss2 @ [n]) ss')" |
|
53 |
by pat_completeness auto |
|
54 |
||
55 |
termination |
|
56 |
sorry |
|
57 |
||
58 |
fun rep :: "Nat \<Rightarrow> RE \<Rightarrow> RE" |
|
59 |
where |
|
60 |
"rep n re = (case n of |
|
61 |
Zer \<Rightarrow> Empty |
|
62 |
| Suc n' \<Rightarrow> Seq re (rep n' re))" |
|
63 |
||
64 |
||
65 |
fun repMax :: "Nat \<Rightarrow> RE \<Rightarrow> RE" |
|
66 |
where |
|
67 |
"repMax n re = (case n of |
|
68 |
Zer \<Rightarrow> Empty |
|
69 |
| Suc n' \<Rightarrow> Or (rep n re) (repMax n' re))" |
|
70 |
||
71 |
||
72 |
fun repInt' :: "Nat \<Rightarrow> Nat \<Rightarrow> RE \<Rightarrow> RE" |
|
73 |
where |
|
74 |
"repInt' n k re = (case k of |
|
75 |
Zer \<Rightarrow> rep n re |
|
76 |
| Suc k' \<Rightarrow> Or (rep n re) (repInt' (Suc n) k' re))" |
|
77 |
||
78 |
||
79 |
fun repInt :: "Nat \<Rightarrow> Nat \<Rightarrow> RE \<Rightarrow> RE" |
|
80 |
where |
|
81 |
"repInt n k re = repInt' n (sub k n) re" |
|
82 |
||
83 |
||
84 |
fun prop_regex :: "Nat * Nat * RE * RE * Sym list \<Rightarrow> bool" |
|
85 |
where |
|
86 |
"prop_regex (n, (k, (p, (q, s)))) = ((accepts (repInt n k (And p q)) s) = (accepts (And (repInt n k p) (repInt n k q)) s))" |
|
87 |
||
88 |
||
89 |
||
90 |
lemma "accepts (repInt n k (And p q)) s --> accepts (And (repInt n k p) (repInt n k q)) s" |
|
91 |
(*nitpick |
|
40924 | 92 |
quickcheck[tester = random, iterations = 10000] |
93 |
quickcheck[tester = predicate_compile_wo_ff] |
|
94 |
quickcheck[tester = predicate_compile_ff_fs] |
|
39188 | 95 |
oops*) |
96 |
oops |
|
97 |
||
98 |
||
39463 | 99 |
setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *} |
39188 | 100 |
|
101 |
setup {* Code_Prolog.map_code_options (K |
|
102 |
{ensure_groundness = true, |
|
39800 | 103 |
limit_globally = NONE, |
39188 | 104 |
limited_types = [(@{typ Sym}, 0), (@{typ "Sym list"}, 2), (@{typ RE}, 6)], |
105 |
limited_predicates = [(["repIntPa"], 2), (["repP"], 2), (["subP"], 0), |
|
39725 | 106 |
(["accepts", "acceptsaux", "seqSplit", "seqSplita", "seqSplitaux", "seqSplitb"], 25)], |
39188 | 107 |
replacing = |
108 |
[(("repP", "limited_repP"), "lim_repIntPa"), |
|
109 |
(("subP", "limited_subP"), "repIntP"), |
|
110 |
(("repIntPa", "limited_repIntPa"), "repIntP"), |
|
111 |
(("accepts", "limited_accepts"), "quickcheck")], |
|
39463 | 112 |
manual_reorder = []}) *} |
39188 | 113 |
|
114 |
text {* Finding the counterexample still seems out of reach for the |
|
115 |
prolog-style generation. *} |
|
116 |
||
117 |
lemma "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s" |
|
40924 | 118 |
quickcheck[tester = random, iterations = 100000, expect = counterexample] |
119 |
(*quickcheck[tester = predicate_compile_wo_ff]*) |
|
120 |
(*quickcheck[tester = predicate_compile_ff_fs, iterations = 1]*) |
|
121 |
(*quickcheck[tester = prolog, iterations = 1, size = 1]*) |
|
39188 | 122 |
oops |
123 |
||
124 |
section {* Given a partial solution *} |
|
125 |
||
126 |
lemma |
|
127 |
(* assumes "n = Zer" |
|
128 |
assumes "k = Suc (Suc Zer)"*) |
|
129 |
assumes "p = Sym N0" |
|
130 |
assumes "q = Seq (Sym N0) (Sym N0)" |
|
131 |
(* assumes "s = [N0, N0]"*) |
|
132 |
shows "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s" |
|
40924 | 133 |
(*quickcheck[tester = predicate_compile_wo_ff, iterations = 1]*) |
134 |
quickcheck[tester = prolog, iterations = 1, size = 1, expect = counterexample] |
|
39188 | 135 |
oops |
136 |
||
137 |
section {* Checking the counterexample *} |
|
138 |
||
139 |
definition a_sol |
|
140 |
where |
|
141 |
"a_sol = (Zer, (Suc (Suc Zer), (Sym N0, (Seq (Sym N0) (Sym N0), [N0, N0]))))" |
|
142 |
||
143 |
lemma |
|
144 |
assumes "n = Zer" |
|
145 |
assumes "k = Suc (Suc Zer)" |
|
146 |
assumes "p = Sym N0" |
|
147 |
assumes "q = Seq (Sym N0) (Sym N0)" |
|
148 |
assumes "s = [N0, N0]" |
|
149 |
shows "accepts (repInt n k (And p q)) s --> accepts (And (repInt n k p) (repInt n k q)) s" |
|
40924 | 150 |
(*quickcheck[tester = predicate_compile_wo_ff]*) |
39188 | 151 |
oops |
152 |
||
153 |
lemma |
|
154 |
assumes "n = Zer" |
|
155 |
assumes "k = Suc (Suc Zer)" |
|
156 |
assumes "p = Sym N0" |
|
157 |
assumes "q = Seq (Sym N0) (Sym N0)" |
|
158 |
assumes "s = [N0, N0]" |
|
159 |
shows "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s" |
|
40924 | 160 |
(*quickcheck[tester = predicate_compile_wo_ff, iterations = 1, expect = counterexample]*) |
161 |
quickcheck[tester = prolog, iterations = 1, size = 1, expect = counterexample] |
|
39188 | 162 |
oops |
163 |
||
164 |
lemma |
|
165 |
assumes "n = Zer" |
|
166 |
assumes "k = Suc (Suc Zer)" |
|
167 |
assumes "p = Sym N0" |
|
168 |
assumes "q = Seq (Sym N0) (Sym N0)" |
|
169 |
assumes "s = [N0, N0]" |
|
170 |
shows "prop_regex (n, (k, (p, (q, s))))" |
|
40924 | 171 |
quickcheck[tester = predicate_compile_wo_ff] |
172 |
quickcheck[tester = prolog, expect = counterexample] |
|
39188 | 173 |
oops |
174 |
||
175 |
lemma "prop_regex a_sol" |
|
40924 | 176 |
quickcheck[tester = random, expect = counterexample] |
177 |
quickcheck[tester = predicate_compile_ff_fs] |
|
39188 | 178 |
oops |
179 |
||
180 |
value [code] "prop_regex a_sol" |
|
181 |
||
182 |
||
183 |
end |