39189
|
1 |
theory Reg_Exp_Example
|
39188
|
2 |
imports Predicate_Compile_Quickcheck Code_Prolog
|
|
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
|
|
92 |
quickcheck[generator = code, iterations = 10000]
|
|
93 |
quickcheck[generator = predicate_compile_wo_ff]
|
|
94 |
quickcheck[generator = predicate_compile_ff_fs]
|
|
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,
|
|
103 |
limited_types = [(@{typ Sym}, 0), (@{typ "Sym list"}, 2), (@{typ RE}, 6)],
|
|
104 |
limited_predicates = [(["repIntPa"], 2), (["repP"], 2), (["subP"], 0),
|
|
105 |
(["accepts", "acceptsaux", "acceptsauxauxa", "acceptsauxaux", "seqSplit", "seqSplita", "seqSplitaux", "seqSplitb"], 25)],
|
|
106 |
replacing =
|
|
107 |
[(("repP", "limited_repP"), "lim_repIntPa"),
|
|
108 |
(("subP", "limited_subP"), "repIntP"),
|
|
109 |
(("repIntPa", "limited_repIntPa"), "repIntP"),
|
|
110 |
(("accepts", "limited_accepts"), "quickcheck")],
|
39463
|
111 |
manual_reorder = []}) *}
|
39188
|
112 |
|
|
113 |
text {* Finding the counterexample still seems out of reach for the
|
|
114 |
prolog-style generation. *}
|
|
115 |
|
|
116 |
lemma "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
|
|
117 |
quickcheck[generator = code, iterations = 100000, expect = counterexample]
|
|
118 |
(*quickcheck[generator = predicate_compile_wo_ff]*)
|
|
119 |
(*quickcheck[generator = predicate_compile_ff_fs, iterations = 1]*)
|
|
120 |
(*quickcheck[generator = prolog, iterations = 1, size = 1]*)
|
|
121 |
oops
|
|
122 |
|
|
123 |
section {* Given a partial solution *}
|
|
124 |
|
|
125 |
lemma
|
|
126 |
(* assumes "n = Zer"
|
|
127 |
assumes "k = Suc (Suc Zer)"*)
|
|
128 |
assumes "p = Sym N0"
|
|
129 |
assumes "q = Seq (Sym N0) (Sym N0)"
|
|
130 |
(* assumes "s = [N0, N0]"*)
|
|
131 |
shows "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
|
|
132 |
(*quickcheck[generator = predicate_compile_wo_ff, iterations = 1]*)
|
|
133 |
quickcheck[generator = prolog, iterations = 1, size = 1, expect = counterexample]
|
|
134 |
oops
|
|
135 |
|
|
136 |
section {* Checking the counterexample *}
|
|
137 |
|
|
138 |
definition a_sol
|
|
139 |
where
|
|
140 |
"a_sol = (Zer, (Suc (Suc Zer), (Sym N0, (Seq (Sym N0) (Sym N0), [N0, N0]))))"
|
|
141 |
|
|
142 |
lemma
|
|
143 |
assumes "n = Zer"
|
|
144 |
assumes "k = Suc (Suc Zer)"
|
|
145 |
assumes "p = Sym N0"
|
|
146 |
assumes "q = Seq (Sym N0) (Sym N0)"
|
|
147 |
assumes "s = [N0, N0]"
|
|
148 |
shows "accepts (repInt n k (And p q)) s --> accepts (And (repInt n k p) (repInt n k q)) s"
|
|
149 |
(*quickcheck[generator = predicate_compile_wo_ff]*)
|
|
150 |
oops
|
|
151 |
|
|
152 |
lemma
|
|
153 |
assumes "n = Zer"
|
|
154 |
assumes "k = Suc (Suc Zer)"
|
|
155 |
assumes "p = Sym N0"
|
|
156 |
assumes "q = Seq (Sym N0) (Sym N0)"
|
|
157 |
assumes "s = [N0, N0]"
|
|
158 |
shows "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
|
39189
|
159 |
(*quickcheck[generator = predicate_compile_wo_ff, iterations = 1, expect = counterexample]*)
|
39188
|
160 |
quickcheck[generator = prolog, iterations = 1, size = 1, expect = counterexample]
|
|
161 |
oops
|
|
162 |
|
|
163 |
lemma
|
|
164 |
assumes "n = Zer"
|
|
165 |
assumes "k = Suc (Suc Zer)"
|
|
166 |
assumes "p = Sym N0"
|
|
167 |
assumes "q = Seq (Sym N0) (Sym N0)"
|
|
168 |
assumes "s = [N0, N0]"
|
|
169 |
shows "prop_regex (n, (k, (p, (q, s))))"
|
39196
|
170 |
quickcheck[generator = predicate_compile_wo_ff]
|
39188
|
171 |
quickcheck[generator = prolog, expect = counterexample]
|
|
172 |
oops
|
|
173 |
|
|
174 |
lemma "prop_regex a_sol"
|
|
175 |
quickcheck[generator = code, expect = counterexample]
|
39196
|
176 |
quickcheck[generator = predicate_compile_ff_fs]
|
39188
|
177 |
oops
|
|
178 |
|
|
179 |
value [code] "prop_regex a_sol"
|
|
180 |
|
|
181 |
|
|
182 |
end
|