259
|
1 |
theory Exercises
|
266
|
2 |
imports Spec "~~/src/HOL/Library/Infinite_Set"
|
258
|
3 |
begin
|
|
4 |
|
260
|
5 |
section {* Some Fun Facts *}
|
258
|
6 |
|
|
7 |
fun
|
259
|
8 |
zeroable :: "rexp \<Rightarrow> bool"
|
258
|
9 |
where
|
259
|
10 |
"zeroable (ZERO) \<longleftrightarrow> True"
|
|
11 |
| "zeroable (ONE) \<longleftrightarrow> False"
|
362
|
12 |
| "zeroable (CH c) \<longleftrightarrow> False"
|
259
|
13 |
| "zeroable (ALT r1 r2) \<longleftrightarrow> zeroable r1 \<and> zeroable r2"
|
|
14 |
| "zeroable (SEQ r1 r2) \<longleftrightarrow> zeroable r1 \<or> zeroable r2"
|
|
15 |
| "zeroable (STAR r) \<longleftrightarrow> False"
|
258
|
16 |
|
|
17 |
lemma zeroable_correctness:
|
259
|
18 |
shows "zeroable r \<longleftrightarrow> L r = {}"
|
260
|
19 |
by(induct r)(auto simp add: Sequ_def)
|
|
20 |
|
258
|
21 |
|
|
22 |
fun
|
259
|
23 |
atmostempty :: "rexp \<Rightarrow> bool"
|
258
|
24 |
where
|
259
|
25 |
"atmostempty (ZERO) \<longleftrightarrow> True"
|
|
26 |
| "atmostempty (ONE) \<longleftrightarrow> True"
|
362
|
27 |
| "atmostempty (CH c) \<longleftrightarrow> False"
|
259
|
28 |
| "atmostempty (ALT r1 r2) \<longleftrightarrow> atmostempty r1 \<and> atmostempty r2"
|
|
29 |
| "atmostempty (SEQ r1 r2) \<longleftrightarrow>
|
|
30 |
zeroable r1 \<or> zeroable r2 \<or> (atmostempty r1 \<and> atmostempty r2)"
|
258
|
31 |
| "atmostempty (STAR r) = atmostempty r"
|
|
32 |
|
641
|
33 |
lemma "atmostempty r \<longrightarrow> (zeroable r \<or> nullable r)"
|
|
34 |
apply(induct r)
|
|
35 |
apply(auto)
|
|
36 |
done
|
283
|
37 |
|
258
|
38 |
fun
|
259
|
39 |
somechars :: "rexp \<Rightarrow> bool"
|
258
|
40 |
where
|
259
|
41 |
"somechars (ZERO) \<longleftrightarrow> False"
|
|
42 |
| "somechars (ONE) \<longleftrightarrow> False"
|
362
|
43 |
| "somechars (CH c) \<longleftrightarrow> True"
|
259
|
44 |
| "somechars (ALT r1 r2) \<longleftrightarrow> somechars r1 \<or> somechars r2"
|
|
45 |
| "somechars (SEQ r1 r2) \<longleftrightarrow>
|
641
|
46 |
(\<not> zeroable r1 \<and> somechars r2) \<or> (\<not> zeroable r2 \<and> somechars r1)"
|
259
|
47 |
| "somechars (STAR r) \<longleftrightarrow> somechars r"
|
258
|
48 |
|
|
49 |
lemma somechars_correctness:
|
259
|
50 |
shows "somechars r \<longleftrightarrow> (\<exists>s. s \<noteq> [] \<and> s \<in> L r)"
|
260
|
51 |
apply(induct r)
|
|
52 |
apply(simp_all add: zeroable_correctness nullable_correctness Sequ_def)
|
|
53 |
using Nil_is_append_conv apply blast
|
|
54 |
apply blast
|
318
|
55 |
apply(auto)
|
|
56 |
by (metis Star_decomp hd_Cons_tl list.distinct(1))
|
258
|
57 |
|
|
58 |
lemma atmostempty_correctness_aux:
|
259
|
59 |
shows "atmostempty r \<longleftrightarrow> \<not> somechars r"
|
258
|
60 |
apply(induct r)
|
|
61 |
apply(simp_all)
|
|
62 |
apply(auto simp add: zeroable_correctness nullable_correctness somechars_correctness)
|
|
63 |
done
|
|
64 |
|
|
65 |
lemma atmostempty_correctness:
|
259
|
66 |
shows "atmostempty r \<longleftrightarrow> L r \<subseteq> {[]}"
|
258
|
67 |
by(auto simp add: atmostempty_correctness_aux somechars_correctness)
|
|
68 |
|
|
69 |
fun
|
283
|
70 |
leastsinglechar :: "rexp \<Rightarrow> bool"
|
|
71 |
where
|
|
72 |
"leastsinglechar (ZERO) \<longleftrightarrow> False"
|
|
73 |
| "leastsinglechar (ONE) \<longleftrightarrow> False"
|
362
|
74 |
| "leastsinglechar (CH c) \<longleftrightarrow> True"
|
283
|
75 |
| "leastsinglechar (ALT r1 r2) \<longleftrightarrow> leastsinglechar r1 \<or> leastsinglechar r2"
|
|
76 |
| "leastsinglechar (SEQ r1 r2) \<longleftrightarrow>
|
|
77 |
(if (zeroable r1 \<or> zeroable r2) then False
|
|
78 |
else ((nullable r1 \<and> leastsinglechar r2) \<or> (nullable r2 \<and> leastsinglechar r1)))"
|
|
79 |
| "leastsinglechar (STAR r) \<longleftrightarrow> leastsinglechar r"
|
|
80 |
|
|
81 |
lemma leastsinglechar_correctness:
|
|
82 |
"leastsinglechar r \<longleftrightarrow> (\<exists>c. [c] \<in> L r)"
|
|
83 |
apply(induct r)
|
|
84 |
apply(simp)
|
|
85 |
apply(simp)
|
|
86 |
apply(simp)
|
|
87 |
prefer 2
|
|
88 |
apply(simp)
|
|
89 |
apply(blast)
|
|
90 |
prefer 2
|
|
91 |
apply(simp)
|
|
92 |
using Star.step Star_decomp apply fastforce
|
|
93 |
apply(simp add: Sequ_def zeroable_correctness nullable_correctness)
|
|
94 |
by (metis append_Nil append_is_Nil_conv butlast_append butlast_snoc)
|
|
95 |
|
|
96 |
fun
|
259
|
97 |
infinitestrings :: "rexp \<Rightarrow> bool"
|
258
|
98 |
where
|
|
99 |
"infinitestrings (ZERO) = False"
|
|
100 |
| "infinitestrings (ONE) = False"
|
362
|
101 |
| "infinitestrings (CH c) = False"
|
259
|
102 |
| "infinitestrings (ALT r1 r2) = (infinitestrings r1 \<or> infinitestrings r2)"
|
|
103 |
| "infinitestrings (SEQ r1 r2) \<longleftrightarrow>
|
|
104 |
(\<not> zeroable r1 \<and> infinitestrings r2) \<or> (\<not> zeroable r2 \<and> infinitestrings r1)"
|
|
105 |
| "infinitestrings (STAR r) = (\<not> atmostempty r)"
|
258
|
106 |
|
282
|
107 |
|
|
108 |
|
|
109 |
|
|
110 |
|
258
|
111 |
lemma Star_atmostempty:
|
259
|
112 |
assumes "A \<subseteq> {[]}"
|
|
113 |
shows "A\<star> \<subseteq> {[]}"
|
318
|
114 |
using assms
|
|
115 |
using Star_decomp concat_eq_Nil_conv empty_iff insert_iff subsetI subset_singletonD
|
|
116 |
apply(auto)
|
|
117 |
proof -
|
|
118 |
fix x :: "char list"
|
|
119 |
assume a1: "x \<in> A\<star>"
|
|
120 |
assume "\<And>c x A. c # x \<in> A\<star> \<Longrightarrow> \<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A\<star>"
|
|
121 |
then have f2: "\<forall>cs C c. \<exists>csa. c # csa \<in> C \<or> c # cs \<notin> C\<star>"
|
|
122 |
by auto
|
|
123 |
obtain cc :: "char list \<Rightarrow> char" and ccs :: "char list \<Rightarrow> char list" where
|
|
124 |
"\<And>cs. cs = [] \<or> cc cs # ccs cs = cs"
|
|
125 |
by (metis (no_types) list.exhaust)
|
|
126 |
then show "x = []"
|
|
127 |
using f2 a1 by (metis assms empty_iff insert_iff list.distinct(1) subset_singletonD)
|
|
128 |
qed
|
|
129 |
|
258
|
130 |
|
|
131 |
lemma Star_empty_string_finite:
|
259
|
132 |
shows "finite ({[]}\<star>)"
|
258
|
133 |
using Star_atmostempty infinite_super by auto
|
|
134 |
|
|
135 |
lemma Star_empty_finite:
|
259
|
136 |
shows "finite ({}\<star>)"
|
258
|
137 |
using Star_atmostempty infinite_super by auto
|
|
138 |
|
|
139 |
lemma Star_concat_replicate:
|
259
|
140 |
assumes "s \<in> A"
|
|
141 |
shows "concat (replicate n s) \<in> A\<star>"
|
258
|
142 |
using assms
|
|
143 |
by (induct n) (auto)
|
|
144 |
|
|
145 |
|
|
146 |
lemma concat_replicate_inj:
|
259
|
147 |
assumes "concat (replicate n s) = concat (replicate m s)" "s \<noteq> []"
|
258
|
148 |
shows "n = m"
|
|
149 |
using assms
|
|
150 |
apply(induct n arbitrary: m)
|
|
151 |
apply(auto)[1]
|
|
152 |
apply(auto)
|
|
153 |
apply(case_tac m)
|
|
154 |
apply(clarify)
|
|
155 |
apply(simp only: replicate.simps concat.simps)
|
|
156 |
apply blast
|
|
157 |
by simp
|
|
158 |
|
|
159 |
lemma A0:
|
259
|
160 |
assumes "finite (A ;; B)" "B \<noteq> {}"
|
258
|
161 |
shows "finite A"
|
259
|
162 |
apply(subgoal_tac "\<exists>s. s \<in> B")
|
258
|
163 |
apply(erule exE)
|
259
|
164 |
apply(subgoal_tac "finite {s1 @ s |s1. s1 \<in> A}")
|
|
165 |
apply(rule_tac f="\<lambda>s1. s1 @ s" in finite_imageD)
|
258
|
166 |
apply(simp add: image_def)
|
|
167 |
apply(smt Collect_cong)
|
|
168 |
apply(simp add: inj_on_def)
|
|
169 |
apply(rule_tac B="A ;; B" in finite_subset)
|
|
170 |
apply(auto simp add: Sequ_def)[1]
|
|
171 |
apply(rule assms(1))
|
|
172 |
using assms(2) by auto
|
|
173 |
|
|
174 |
lemma A1:
|
259
|
175 |
assumes "finite (A ;; B)" "A \<noteq> {}"
|
258
|
176 |
shows "finite B"
|
259
|
177 |
apply(subgoal_tac "\<exists>s. s \<in> A")
|
258
|
178 |
apply(erule exE)
|
259
|
179 |
apply(subgoal_tac "finite {s @ s1 |s1. s1 \<in> B}")
|
|
180 |
apply(rule_tac f="\<lambda>s1. s @ s1" in finite_imageD)
|
258
|
181 |
apply(simp add: image_def)
|
|
182 |
apply(smt Collect_cong)
|
|
183 |
apply(simp add: inj_on_def)
|
|
184 |
apply(rule_tac B="A ;; B" in finite_subset)
|
|
185 |
apply(auto simp add: Sequ_def)[1]
|
|
186 |
apply(rule assms(1))
|
|
187 |
using assms(2) by auto
|
|
188 |
|
|
189 |
lemma Sequ_Prod_finite:
|
259
|
190 |
assumes "A \<noteq> {}" "B \<noteq> {}"
|
|
191 |
shows "finite (A ;; B) \<longleftrightarrow> (finite (A \<times> B))"
|
258
|
192 |
apply(rule iffI)
|
|
193 |
apply(rule finite_cartesian_product)
|
|
194 |
apply(erule A0)
|
|
195 |
apply(rule assms(2))
|
|
196 |
apply(erule A1)
|
|
197 |
apply(rule assms(1))
|
|
198 |
apply(simp add: Sequ_def)
|
|
199 |
apply(rule finite_image_set2)
|
|
200 |
apply(drule finite_cartesian_productD1)
|
|
201 |
apply(rule assms(2))
|
|
202 |
apply(simp)
|
|
203 |
apply(drule finite_cartesian_productD2)
|
|
204 |
apply(rule assms(1))
|
|
205 |
apply(simp)
|
|
206 |
done
|
|
207 |
|
|
208 |
|
|
209 |
lemma Star_non_empty_string_infinite:
|
259
|
210 |
assumes "s \<in> A" " s \<noteq> []"
|
|
211 |
shows "infinite (A\<star>)"
|
258
|
212 |
proof -
|
259
|
213 |
have "inj (\<lambda>n. concat (replicate n s))"
|
258
|
214 |
using assms(2) concat_replicate_inj
|
|
215 |
by(auto simp add: inj_on_def)
|
|
216 |
moreover
|
|
217 |
have "infinite (UNIV::nat set)" by simp
|
|
218 |
ultimately
|
259
|
219 |
have "infinite ((\<lambda>n. concat (replicate n s)) ` UNIV)"
|
258
|
220 |
by (simp add: range_inj_infinite)
|
|
221 |
moreover
|
259
|
222 |
have "((\<lambda>n. concat (replicate n s)) ` UNIV) \<subseteq> (A\<star>)"
|
258
|
223 |
using Star_concat_replicate assms(1) by auto
|
259
|
224 |
ultimately show "infinite (A\<star>)"
|
258
|
225 |
using infinite_super by auto
|
|
226 |
qed
|
|
227 |
|
|
228 |
lemma infinitestrings_correctness:
|
259
|
229 |
shows "infinitestrings r \<longleftrightarrow> infinite (L r)"
|
258
|
230 |
apply(induct r)
|
|
231 |
apply(simp_all)
|
|
232 |
apply(simp add: zeroable_correctness)
|
|
233 |
apply(rule iffI)
|
|
234 |
apply(erule disjE)
|
|
235 |
apply(subst Sequ_Prod_finite)
|
|
236 |
apply(auto)[2]
|
|
237 |
using finite_cartesian_productD2 apply blast
|
|
238 |
apply(subst Sequ_Prod_finite)
|
|
239 |
apply(auto)[2]
|
|
240 |
using finite_cartesian_productD1 apply blast
|
259
|
241 |
apply(subgoal_tac "L r1 \<noteq> {} \<and> L r2 \<noteq> {}")
|
258
|
242 |
prefer 2
|
|
243 |
apply(auto simp add: Sequ_def)[1]
|
|
244 |
apply(subst (asm) Sequ_Prod_finite)
|
|
245 |
apply(auto)[2]
|
|
246 |
apply(auto)[1]
|
|
247 |
apply(simp add: atmostempty_correctness)
|
|
248 |
apply(rule iffI)
|
|
249 |
apply (metis Star_empty_finite Star_empty_string_finite subset_singletonD)
|
|
250 |
using Star_non_empty_string_infinite apply blast
|
|
251 |
done
|
|
252 |
|
260
|
253 |
unused_thms
|
258
|
254 |
|
|
255 |
end |