author | urbanc |
Thu, 01 Sep 2011 23:18:34 +0000 | |
changeset 222 | 191769fc68c3 |
parent 221 | 68e28debe995 |
child 223 | 5f7b7ad498dd |
permissions | -rw-r--r-- |
220 | 1 |
theory Closure2 |
221 | 2 |
imports |
3 |
Closures |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
4 |
Higman |
221 | 5 |
(* "~~/src/HOL/Proofs/Extraction/Higman" *) |
220 | 6 |
begin |
7 |
||
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
8 |
notation |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
9 |
emb ("_ \<preceq> _") |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
10 |
|
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
11 |
declare emb0 [intro] |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
12 |
declare emb1 [intro] |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
13 |
declare emb2 [intro] |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
14 |
|
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
15 |
lemma letter_UNIV: |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
16 |
shows "UNIV = {A, B::letter}" |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
17 |
apply(auto) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
18 |
apply(case_tac x) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
19 |
apply(auto) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
20 |
done |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
21 |
|
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
22 |
instance letter :: finite |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
23 |
apply(default) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
24 |
apply(simp add: letter_UNIV) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
25 |
done |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
26 |
|
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
27 |
hide_const A |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
28 |
hide_const B |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
29 |
|
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
30 |
(* |
221 | 31 |
inductive |
32 |
emb :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<preceq> _") |
|
220 | 33 |
where |
221 | 34 |
emb0 [intro]: "emb [] y" |
35 |
| emb1 [intro]: "emb x y \<Longrightarrow> emb x (c # y)" |
|
36 |
| emb2 [intro]: "emb x y \<Longrightarrow> emb (c # x) (c # y)" |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
37 |
*) |
220 | 38 |
|
221 | 39 |
lemma emb_refl [intro]: |
220 | 40 |
shows "x \<preceq> x" |
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
41 |
by (induct x) (auto) |
220 | 42 |
|
221 | 43 |
lemma emb_right [intro]: |
220 | 44 |
assumes a: "x \<preceq> y" |
45 |
shows "x \<preceq> y @ y'" |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
46 |
using a |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
47 |
by (induct arbitrary: y') (auto) |
220 | 48 |
|
221 | 49 |
lemma emb_left [intro]: |
220 | 50 |
assumes a: "x \<preceq> y" |
51 |
shows "x \<preceq> y' @ y" |
|
221 | 52 |
using a by (induct y') (auto) |
220 | 53 |
|
221 | 54 |
lemma emb_appendI [intro]: |
220 | 55 |
assumes a: "x \<preceq> x'" |
56 |
and b: "y \<preceq> y'" |
|
57 |
shows "x @ y \<preceq> x' @ y'" |
|
221 | 58 |
using a b by (induct) (auto) |
59 |
||
60 |
lemma emb_cons_leftD: |
|
61 |
assumes "a # x \<preceq> y" |
|
62 |
shows "\<exists>y1 y2. y = y1 @ [a] @ y2 \<and> x \<preceq> y2" |
|
63 |
using assms |
|
64 |
apply(induct x\<equiv>"a # x" y\<equiv>"y" arbitrary: a x y) |
|
65 |
apply(auto) |
|
66 |
apply(metis append_Cons) |
|
67 |
done |
|
68 |
||
69 |
lemma emb_append_leftD: |
|
70 |
assumes "x1 @ x2 \<preceq> y" |
|
71 |
shows "\<exists>y1 y2. y = y1 @ y2 \<and> x1 \<preceq> y1 \<and> x2 \<preceq> y2" |
|
72 |
using assms |
|
73 |
apply(induct x1 arbitrary: x2 y) |
|
74 |
apply(auto) |
|
75 |
apply(drule emb_cons_leftD) |
|
76 |
apply(auto) |
|
77 |
apply(drule_tac x="x2" in meta_spec) |
|
78 |
apply(drule_tac x="y2" in meta_spec) |
|
79 |
apply(auto) |
|
80 |
apply(rule_tac x="y1 @ (a # y1a)" in exI) |
|
81 |
apply(rule_tac x="y2a" in exI) |
|
82 |
apply(auto) |
|
220 | 83 |
done |
84 |
||
221 | 85 |
lemma emb_trans: |
86 |
assumes a: "x1 \<preceq> x2" |
|
87 |
and b: "x2 \<preceq> x3" |
|
88 |
shows "x1 \<preceq> x3" |
|
89 |
using a b |
|
90 |
apply(induct arbitrary: x3) |
|
91 |
apply(metis emb0) |
|
92 |
apply(metis emb_cons_leftD emb_left) |
|
93 |
apply(drule_tac emb_cons_leftD) |
|
94 |
apply(auto) |
|
220 | 95 |
done |
96 |
||
221 | 97 |
lemma emb_strict_length: |
98 |
assumes a: "x \<preceq> y" "x \<noteq> y" |
|
99 |
shows "length x < length y" |
|
100 |
using a |
|
101 |
by (induct) (auto simp add: less_Suc_eq) |
|
102 |
||
103 |
lemma emb_antisym: |
|
104 |
assumes a: "x \<preceq> y" "y \<preceq> x" |
|
105 |
shows "x = y" |
|
106 |
using a emb_strict_length |
|
107 |
by (metis not_less_iff_gr_or_eq) |
|
108 |
||
109 |
lemma emb_wf: |
|
110 |
shows "wf {(x, y). x \<preceq> y \<and> x \<noteq> y}" |
|
111 |
proof - |
|
112 |
have "wf (measure length)" by simp |
|
113 |
moreover |
|
114 |
have "{(x, y). x \<preceq> y \<and> x \<noteq> y} \<subseteq> measure length" |
|
115 |
unfolding measure_def by (auto simp add: emb_strict_length) |
|
116 |
ultimately |
|
117 |
show "wf {(x, y). x \<preceq> y \<and> x \<noteq> y}" by (rule wf_subset) |
|
118 |
qed |
|
119 |
||
120 |
subsection {* Sub- and Supsequences *} |
|
220 | 121 |
|
122 |
definition |
|
123 |
"SUBSEQ A \<equiv> {x. \<exists>y \<in> A. x \<preceq> y}" |
|
124 |
||
125 |
definition |
|
221 | 126 |
"SUPSEQ A \<equiv> {x. \<exists>y \<in> A. y \<preceq> x}" |
220 | 127 |
|
221 | 128 |
lemma SUPSEQ_simps [simp]: |
129 |
shows "SUPSEQ {} = {}" |
|
130 |
and "SUPSEQ {[]} = UNIV" |
|
131 |
unfolding SUPSEQ_def by auto |
|
220 | 132 |
|
221 | 133 |
lemma SUPSEQ_atom [simp]: |
134 |
shows "SUPSEQ {[a]} = UNIV \<cdot> {[a]} \<cdot> UNIV" |
|
135 |
unfolding SUPSEQ_def conc_def |
|
136 |
by (auto) (metis append_Cons emb_cons_leftD) |
|
220 | 137 |
|
221 | 138 |
lemma SUPSEQ_union [simp]: |
139 |
shows "SUPSEQ (A \<union> B) = SUPSEQ A \<union> SUPSEQ B" |
|
140 |
unfolding SUPSEQ_def by auto |
|
220 | 141 |
|
221 | 142 |
lemma SUPSEQ_conc [simp]: |
143 |
shows "SUPSEQ (A \<cdot> B) = SUPSEQ A \<cdot> SUPSEQ B" |
|
144 |
unfolding SUPSEQ_def conc_def |
|
145 |
by (auto) (metis emb_append_leftD) |
|
220 | 146 |
|
221 | 147 |
lemma SUPSEQ_star [simp]: |
148 |
shows "SUPSEQ (A\<star>) = UNIV" |
|
220 | 149 |
apply(subst star_unfold_left) |
221 | 150 |
apply(simp only: SUPSEQ_union) |
220 | 151 |
apply(simp) |
152 |
done |
|
153 |
||
154 |
definition |
|
155 |
Allreg :: "'a::finite rexp" |
|
156 |
where |
|
157 |
"Allreg \<equiv> \<Uplus>(Atom ` UNIV)" |
|
158 |
||
159 |
lemma Allreg_lang [simp]: |
|
221 | 160 |
shows "lang Allreg = (\<Union>a. {[a]})" |
161 |
unfolding Allreg_def by auto |
|
220 | 162 |
|
163 |
lemma [simp]: |
|
221 | 164 |
shows "(\<Union>a. {[a]})\<star> = UNIV" |
220 | 165 |
apply(auto) |
166 |
apply(induct_tac x rule: list.induct) |
|
167 |
apply(auto) |
|
168 |
apply(subgoal_tac "[a] @ list \<in> (\<Union>a. {[a]})\<star>") |
|
169 |
apply(simp) |
|
170 |
apply(rule append_in_starI) |
|
171 |
apply(auto) |
|
172 |
done |
|
173 |
||
174 |
lemma Star_Allreg_lang [simp]: |
|
221 | 175 |
shows "lang (Star Allreg) = UNIV" |
176 |
by simp |
|
220 | 177 |
|
221 | 178 |
fun |
179 |
UP :: "'a::finite rexp \<Rightarrow> 'a rexp" |
|
180 |
where |
|
181 |
"UP (Zero) = Zero" |
|
182 |
| "UP (One) = Star Allreg" |
|
183 |
| "UP (Atom c) = Times (Star Allreg) (Times (Atom c) (Star Allreg))" |
|
184 |
| "UP (Plus r1 r2) = Plus (UP r1) (UP r2)" |
|
185 |
| "UP (Times r1 r2) = Times (UP r1) (UP r2)" |
|
186 |
| "UP (Star r) = Star Allreg" |
|
220 | 187 |
|
221 | 188 |
lemma lang_UP: |
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
189 |
fixes r::"letter rexp" |
221 | 190 |
shows "lang (UP r) = SUPSEQ (lang r)" |
191 |
by (induct r) (simp_all) |
|
192 |
||
193 |
lemma regular_SUPSEQ: |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
194 |
fixes A::"letter lang" |
220 | 195 |
assumes "regular A" |
196 |
shows "regular (SUPSEQ A)" |
|
197 |
proof - |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
198 |
from assms obtain r::"letter rexp" where "lang r = A" by auto |
221 | 199 |
then have "lang (UP r) = SUPSEQ A" by (simp add: lang_UP) |
200 |
then show "regular (SUPSEQ A)" by auto |
|
220 | 201 |
qed |
221 | 202 |
|
203 |
lemma SUPSEQ_subset: |
|
204 |
shows "A \<subseteq> SUPSEQ A" |
|
205 |
unfolding SUPSEQ_def by auto |
|
206 |
||
207 |
lemma w3: |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
208 |
fixes T A::"letter lang" |
221 | 209 |
assumes eq: "T = - (SUBSEQ A)" |
210 |
shows "T = SUPSEQ T" |
|
211 |
apply(rule) |
|
212 |
apply(rule SUPSEQ_subset) |
|
213 |
apply(rule ccontr) |
|
214 |
apply(auto) |
|
215 |
apply(rule ccontr) |
|
216 |
apply(subgoal_tac "x \<in> SUBSEQ A") |
|
217 |
prefer 2 |
|
218 |
apply(subst (asm) (2) eq) |
|
219 |
apply(simp) |
|
220 |
apply(simp add: SUPSEQ_def) |
|
221 |
apply(erule bexE) |
|
222 |
apply(subgoal_tac "y \<in> SUBSEQ A") |
|
223 |
prefer 2 |
|
224 |
apply(simp add: SUBSEQ_def) |
|
225 |
apply(erule bexE) |
|
226 |
apply(rule_tac x="xa" in bexI) |
|
227 |
apply(rule emb_trans) |
|
228 |
apply(assumption) |
|
229 |
apply(assumption) |
|
230 |
apply(assumption) |
|
231 |
apply(subst (asm) (2) eq) |
|
232 |
apply(simp) |
|
233 |
done |
|
234 |
||
235 |
lemma w4: |
|
236 |
shows "- (SUBSEQ A) = SUPSEQ (- (SUBSEQ A))" |
|
237 |
by (rule w3) (simp) |
|
238 |
||
239 |
definition |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
240 |
minimal_in :: "letter list \<Rightarrow> letter lang \<Rightarrow> bool" |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
241 |
where |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
242 |
"minimal_in x A \<equiv> \<forall>y \<in> A. y \<preceq> x \<longrightarrow> y = x" |
221 | 243 |
|
244 |
lemma minimal_in2: |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
245 |
shows "minimal_in x A = (\<forall>y \<in> A. y \<preceq> x \<longrightarrow> x \<preceq> y)" |
221 | 246 |
by (auto simp add: minimal_in_def intro: emb_antisym) |
247 |
||
248 |
lemma higman: |
|
249 |
assumes "\<forall>x \<in> A. \<forall>y \<in> A. x \<noteq> y \<longrightarrow> \<not>(x \<preceq> y) \<and> \<not>(y \<preceq> x)" |
|
250 |
shows "finite A" |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
251 |
apply(rule ccontr) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
252 |
apply(simp add: infinite_iff_countable_subset) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
253 |
apply(auto) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
254 |
apply(insert higman_idx) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
255 |
apply(drule_tac x="f" in meta_spec) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
256 |
apply(auto) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
257 |
using assms |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
258 |
apply - |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
259 |
apply(drule_tac x="f i" in bspec) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
260 |
apply(auto)[1] |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
261 |
apply(drule_tac x="f j" in bspec) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
262 |
apply(auto)[1] |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
263 |
apply(drule mp) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
264 |
apply(simp add: inj_on_def) |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
265 |
apply(auto)[1] |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
266 |
apply(auto)[1] |
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
267 |
done |
221 | 268 |
|
269 |
lemma minimal: |
|
270 |
assumes "minimal_in x A" "minimal_in y A" |
|
271 |
and "x \<noteq> y" "x \<in> A" "y \<in> A" |
|
272 |
shows "\<not>(x \<preceq> y) \<and> \<not>(y \<preceq> x)" |
|
273 |
using assms unfolding minimal_in_def |
|
274 |
by auto |
|
220 | 275 |
|
221 | 276 |
lemma main_lemma: |
277 |
"\<exists>M. M \<subseteq> A \<and> finite M \<and> SUPSEQ A = SUPSEQ M" |
|
278 |
proof - |
|
279 |
def M \<equiv> "{x \<in> A. minimal_in x A}" |
|
280 |
have "M \<subseteq> A" unfolding M_def minimal_in_def by auto |
|
281 |
moreover |
|
282 |
have "finite M" |
|
283 |
apply(rule higman) |
|
284 |
unfolding M_def minimal_in_def |
|
285 |
by auto |
|
286 |
moreover |
|
287 |
have "SUPSEQ A \<subseteq> SUPSEQ M" |
|
288 |
apply(rule) |
|
289 |
apply(simp only: SUPSEQ_def) |
|
290 |
apply(auto)[1] |
|
291 |
using emb_wf |
|
292 |
apply(erule_tac Q="{y' \<in> A. y' \<preceq> x}" and x="y" in wfE_min) |
|
293 |
apply(simp) |
|
294 |
apply(simp) |
|
295 |
apply(rule_tac x="z" in bexI) |
|
296 |
apply(simp) |
|
297 |
apply(simp add: M_def) |
|
298 |
apply(simp add: minimal_in2) |
|
299 |
apply(rule ballI) |
|
300 |
apply(rule impI) |
|
301 |
apply(drule_tac x="ya" in meta_spec) |
|
302 |
apply(simp) |
|
303 |
apply(case_tac "ya = z") |
|
304 |
apply(auto)[1] |
|
305 |
apply(simp) |
|
306 |
by (metis emb_trans) |
|
307 |
moreover |
|
308 |
have "SUPSEQ M \<subseteq> SUPSEQ A" |
|
309 |
by (auto simp add: SUPSEQ_def M_def) |
|
310 |
ultimately |
|
311 |
show "\<exists>M. M \<subseteq> A \<and> finite M \<and> SUPSEQ A = SUPSEQ M" by blast |
|
312 |
qed |
|
313 |
||
314 |
lemma closure_SUPSEQ: |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
315 |
fixes A::"letter lang" |
221 | 316 |
shows "regular (SUPSEQ A)" |
317 |
proof - |
|
318 |
obtain M where a: "finite M" and b: "SUPSEQ A = SUPSEQ M" |
|
319 |
using main_lemma by blast |
|
320 |
have "regular M" using a by (rule finite_regular) |
|
321 |
then have "regular (SUPSEQ M)" by (rule regular_SUPSEQ) |
|
322 |
then show "regular (SUPSEQ A)" using b by simp |
|
323 |
qed |
|
324 |
||
325 |
lemma closure_SUBSEQ: |
|
222
191769fc68c3
included Higman's lemma from the Isabelle repository
urbanc
parents:
221
diff
changeset
|
326 |
fixes A::"letter lang" |
221 | 327 |
shows "regular (SUBSEQ A)" |
328 |
proof - |
|
329 |
have "regular (SUPSEQ (- SUBSEQ A))" by (rule closure_SUPSEQ) |
|
330 |
then have "regular (- SUBSEQ A)" |
|
331 |
apply(subst w4) |
|
332 |
apply(assumption) |
|
333 |
done |
|
334 |
then have "regular (- (- (SUBSEQ A)))" by (rule closure_complement) |
|
335 |
then show "regular (SUBSEQ A)" by simp |
|
336 |
qed |
|
337 |
||
220 | 338 |
|
339 |
||
340 |
end |