|
1 (* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Gram_Lang.thy |
|
2 Author: Andrei Popescu, TU Muenchen |
|
3 Copyright 2012 |
|
4 |
|
5 Language of a grammar. |
|
6 *) |
|
7 |
|
8 header {* Language of a Grammar *} |
|
9 |
|
10 theory Gram_Lang |
|
11 imports Tree |
|
12 begin |
|
13 |
|
14 |
|
15 consts P :: "(N \<times> (T + N) set) set" |
|
16 axiomatization where |
|
17 finite_N: "finite (UNIV::N set)" |
|
18 and finite_in_P: "\<And> n tns. (n,tns) \<in> P \<longrightarrow> finite tns" |
|
19 and used: "\<And> n. \<exists> tns. (n,tns) \<in> P" |
|
20 |
|
21 |
|
22 subsection{* Tree basics: frontier, interior, etc. *} |
|
23 |
|
24 lemma Tree_cong: |
|
25 assumes "root tr = root tr'" and "cont tr = cont tr'" |
|
26 shows "tr = tr'" |
|
27 by (metis Node_root_cont assms) |
|
28 |
|
29 inductive finiteT where |
|
30 Node: "\<lbrakk>finite as; (finiteT^#) as\<rbrakk> \<Longrightarrow> finiteT (Node a as)" |
|
31 monos lift_mono |
|
32 |
|
33 lemma finiteT_induct[consumes 1, case_names Node, induct pred: finiteT]: |
|
34 assumes 1: "finiteT tr" |
|
35 and IH: "\<And>as n. \<lbrakk>finite as; (\<phi>^#) as\<rbrakk> \<Longrightarrow> \<phi> (Node n as)" |
|
36 shows "\<phi> tr" |
|
37 using 1 apply(induct rule: finiteT.induct) |
|
38 apply(rule IH) apply assumption apply(elim mono_lift) by simp |
|
39 |
|
40 |
|
41 (* Frontier *) |
|
42 |
|
43 inductive inFr :: "N set \<Rightarrow> Tree \<Rightarrow> T \<Rightarrow> bool" where |
|
44 Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr ns tr t" |
|
45 | |
|
46 Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inFr ns tr1 t\<rbrakk> \<Longrightarrow> inFr ns tr t" |
|
47 |
|
48 definition "Fr ns tr \<equiv> {t. inFr ns tr t}" |
|
49 |
|
50 lemma inFr_root_in: "inFr ns tr t \<Longrightarrow> root tr \<in> ns" |
|
51 by (metis inFr.simps) |
|
52 |
|
53 lemma inFr_mono: |
|
54 assumes "inFr ns tr t" and "ns \<subseteq> ns'" |
|
55 shows "inFr ns' tr t" |
|
56 using assms apply(induct arbitrary: ns' rule: inFr.induct) |
|
57 using Base Ind by (metis inFr.simps set_mp)+ |
|
58 |
|
59 lemma inFr_Ind_minus: |
|
60 assumes "inFr ns1 tr1 t" and "Inr tr1 \<in> cont tr" |
|
61 shows "inFr (insert (root tr) ns1) tr t" |
|
62 using assms apply(induct rule: inFr.induct) |
|
63 apply (metis inFr.simps insert_iff) |
|
64 by (metis inFr.simps inFr_mono insertI1 subset_insertI) |
|
65 |
|
66 (* alternative definition *) |
|
67 inductive inFr2 :: "N set \<Rightarrow> Tree \<Rightarrow> T \<Rightarrow> bool" where |
|
68 Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr2 ns tr t" |
|
69 | |
|
70 Ind: "\<lbrakk>Inr tr1 \<in> cont tr; inFr2 ns1 tr1 t\<rbrakk> |
|
71 \<Longrightarrow> inFr2 (insert (root tr) ns1) tr t" |
|
72 |
|
73 lemma inFr2_root_in: "inFr2 ns tr t \<Longrightarrow> root tr \<in> ns" |
|
74 apply(induct rule: inFr2.induct) by auto |
|
75 |
|
76 lemma inFr2_mono: |
|
77 assumes "inFr2 ns tr t" and "ns \<subseteq> ns'" |
|
78 shows "inFr2 ns' tr t" |
|
79 using assms apply(induct arbitrary: ns' rule: inFr2.induct) |
|
80 using Base Ind |
|
81 apply (metis subsetD) by (metis inFr2.simps insert_absorb insert_subset) |
|
82 |
|
83 lemma inFr2_Ind: |
|
84 assumes "inFr2 ns tr1 t" and "root tr \<in> ns" and "Inr tr1 \<in> cont tr" |
|
85 shows "inFr2 ns tr t" |
|
86 using assms apply(induct rule: inFr2.induct) |
|
87 apply (metis inFr2.simps insert_absorb) |
|
88 by (metis inFr2.simps insert_absorb) |
|
89 |
|
90 lemma inFr_inFr2: |
|
91 "inFr = inFr2" |
|
92 apply (rule ext)+ apply(safe) |
|
93 apply(erule inFr.induct) |
|
94 apply (metis (lifting) inFr2.Base) |
|
95 apply (metis (lifting) inFr2_Ind) |
|
96 apply(erule inFr2.induct) |
|
97 apply (metis (lifting) inFr.Base) |
|
98 apply (metis (lifting) inFr_Ind_minus) |
|
99 done |
|
100 |
|
101 lemma not_root_inFr: |
|
102 assumes "root tr \<notin> ns" |
|
103 shows "\<not> inFr ns tr t" |
|
104 by (metis assms inFr_root_in) |
|
105 |
|
106 theorem not_root_Fr: |
|
107 assumes "root tr \<notin> ns" |
|
108 shows "Fr ns tr = {}" |
|
109 using not_root_inFr[OF assms] unfolding Fr_def by auto |
|
110 |
|
111 |
|
112 (* Interior *) |
|
113 |
|
114 inductive inItr :: "N set \<Rightarrow> Tree \<Rightarrow> N \<Rightarrow> bool" where |
|
115 Base: "root tr \<in> ns \<Longrightarrow> inItr ns tr (root tr)" |
|
116 | |
|
117 Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inItr ns tr1 n\<rbrakk> \<Longrightarrow> inItr ns tr n" |
|
118 |
|
119 definition "Itr ns tr \<equiv> {n. inItr ns tr n}" |
|
120 |
|
121 lemma inItr_root_in: "inItr ns tr n \<Longrightarrow> root tr \<in> ns" |
|
122 by (metis inItr.simps) |
|
123 |
|
124 lemma inItr_mono: |
|
125 assumes "inItr ns tr n" and "ns \<subseteq> ns'" |
|
126 shows "inItr ns' tr n" |
|
127 using assms apply(induct arbitrary: ns' rule: inItr.induct) |
|
128 using Base Ind by (metis inItr.simps set_mp)+ |
|
129 |
|
130 |
|
131 (* The subtree relation *) |
|
132 |
|
133 inductive subtr where |
|
134 Refl: "root tr \<in> ns \<Longrightarrow> subtr ns tr tr" |
|
135 | |
|
136 Step: "\<lbrakk>root tr3 \<in> ns; subtr ns tr1 tr2; Inr tr2 \<in> cont tr3\<rbrakk> \<Longrightarrow> subtr ns tr1 tr3" |
|
137 |
|
138 lemma subtr_rootL_in: |
|
139 assumes "subtr ns tr1 tr2" |
|
140 shows "root tr1 \<in> ns" |
|
141 using assms apply(induct rule: subtr.induct) by auto |
|
142 |
|
143 lemma subtr_rootR_in: |
|
144 assumes "subtr ns tr1 tr2" |
|
145 shows "root tr2 \<in> ns" |
|
146 using assms apply(induct rule: subtr.induct) by auto |
|
147 |
|
148 lemmas subtr_roots_in = subtr_rootL_in subtr_rootR_in |
|
149 |
|
150 lemma subtr_mono: |
|
151 assumes "subtr ns tr1 tr2" and "ns \<subseteq> ns'" |
|
152 shows "subtr ns' tr1 tr2" |
|
153 using assms apply(induct arbitrary: ns' rule: subtr.induct) |
|
154 using Refl Step by (metis subtr.simps set_mp)+ |
|
155 |
|
156 lemma subtr_trans_Un: |
|
157 assumes "subtr ns12 tr1 tr2" and "subtr ns23 tr2 tr3" |
|
158 shows "subtr (ns12 \<union> ns23) tr1 tr3" |
|
159 proof- |
|
160 have "subtr ns23 tr2 tr3 \<Longrightarrow> |
|
161 (\<forall> ns12 tr1. subtr ns12 tr1 tr2 \<longrightarrow> subtr (ns12 \<union> ns23) tr1 tr3)" |
|
162 apply(induct rule: subtr.induct, safe) |
|
163 apply (metis subtr_mono sup_commute sup_ge2) |
|
164 by (metis (lifting) Step UnI2) |
|
165 thus ?thesis using assms by auto |
|
166 qed |
|
167 |
|
168 lemma subtr_trans: |
|
169 assumes "subtr ns tr1 tr2" and "subtr ns tr2 tr3" |
|
170 shows "subtr ns tr1 tr3" |
|
171 using subtr_trans_Un[OF assms] by simp |
|
172 |
|
173 lemma subtr_StepL: |
|
174 assumes r: "root tr1 \<in> ns" and tr12: "Inr tr1 \<in> cont tr2" and s: "subtr ns tr2 tr3" |
|
175 shows "subtr ns tr1 tr3" |
|
176 apply(rule subtr_trans[OF _ s]) apply(rule Step[of tr2 ns tr1 tr1]) |
|
177 by (metis assms subtr_rootL_in Refl)+ |
|
178 |
|
179 (* alternative definition: *) |
|
180 inductive subtr2 where |
|
181 Refl: "root tr \<in> ns \<Longrightarrow> subtr2 ns tr tr" |
|
182 | |
|
183 Step: "\<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr2 ns tr2 tr3\<rbrakk> \<Longrightarrow> subtr2 ns tr1 tr3" |
|
184 |
|
185 lemma subtr2_rootL_in: |
|
186 assumes "subtr2 ns tr1 tr2" |
|
187 shows "root tr1 \<in> ns" |
|
188 using assms apply(induct rule: subtr2.induct) by auto |
|
189 |
|
190 lemma subtr2_rootR_in: |
|
191 assumes "subtr2 ns tr1 tr2" |
|
192 shows "root tr2 \<in> ns" |
|
193 using assms apply(induct rule: subtr2.induct) by auto |
|
194 |
|
195 lemmas subtr2_roots_in = subtr2_rootL_in subtr2_rootR_in |
|
196 |
|
197 lemma subtr2_mono: |
|
198 assumes "subtr2 ns tr1 tr2" and "ns \<subseteq> ns'" |
|
199 shows "subtr2 ns' tr1 tr2" |
|
200 using assms apply(induct arbitrary: ns' rule: subtr2.induct) |
|
201 using Refl Step by (metis subtr2.simps set_mp)+ |
|
202 |
|
203 lemma subtr2_trans_Un: |
|
204 assumes "subtr2 ns12 tr1 tr2" and "subtr2 ns23 tr2 tr3" |
|
205 shows "subtr2 (ns12 \<union> ns23) tr1 tr3" |
|
206 proof- |
|
207 have "subtr2 ns12 tr1 tr2 \<Longrightarrow> |
|
208 (\<forall> ns23 tr3. subtr2 ns23 tr2 tr3 \<longrightarrow> subtr2 (ns12 \<union> ns23) tr1 tr3)" |
|
209 apply(induct rule: subtr2.induct, safe) |
|
210 apply (metis subtr2_mono sup_commute sup_ge2) |
|
211 by (metis Un_iff subtr2.simps) |
|
212 thus ?thesis using assms by auto |
|
213 qed |
|
214 |
|
215 lemma subtr2_trans: |
|
216 assumes "subtr2 ns tr1 tr2" and "subtr2 ns tr2 tr3" |
|
217 shows "subtr2 ns tr1 tr3" |
|
218 using subtr2_trans_Un[OF assms] by simp |
|
219 |
|
220 lemma subtr2_StepR: |
|
221 assumes r: "root tr3 \<in> ns" and tr23: "Inr tr2 \<in> cont tr3" and s: "subtr2 ns tr1 tr2" |
|
222 shows "subtr2 ns tr1 tr3" |
|
223 apply(rule subtr2_trans[OF s]) apply(rule Step[of _ _ tr3]) |
|
224 by (metis assms subtr2_rootR_in Refl)+ |
|
225 |
|
226 lemma subtr_subtr2: |
|
227 "subtr = subtr2" |
|
228 apply (rule ext)+ apply(safe) |
|
229 apply(erule subtr.induct) |
|
230 apply (metis (lifting) subtr2.Refl) |
|
231 apply (metis (lifting) subtr2_StepR) |
|
232 apply(erule subtr2.induct) |
|
233 apply (metis (lifting) subtr.Refl) |
|
234 apply (metis (lifting) subtr_StepL) |
|
235 done |
|
236 |
|
237 lemma subtr_inductL[consumes 1, case_names Refl Step]: |
|
238 assumes s: "subtr ns tr1 tr2" and Refl: "\<And>ns tr. \<phi> ns tr tr" |
|
239 and Step: |
|
240 "\<And>ns tr1 tr2 tr3. |
|
241 \<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr ns tr2 tr3; \<phi> ns tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> ns tr1 tr3" |
|
242 shows "\<phi> ns tr1 tr2" |
|
243 using s unfolding subtr_subtr2 apply(rule subtr2.induct) |
|
244 using Refl Step unfolding subtr_subtr2 by auto |
|
245 |
|
246 lemma subtr_UNIV_inductL[consumes 1, case_names Refl Step]: |
|
247 assumes s: "subtr UNIV tr1 tr2" and Refl: "\<And>tr. \<phi> tr tr" |
|
248 and Step: |
|
249 "\<And>tr1 tr2 tr3. |
|
250 \<lbrakk>Inr tr1 \<in> cont tr2; subtr UNIV tr2 tr3; \<phi> tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> tr1 tr3" |
|
251 shows "\<phi> tr1 tr2" |
|
252 using s apply(induct rule: subtr_inductL) |
|
253 apply(rule Refl) using Step subtr_mono by (metis subset_UNIV) |
|
254 |
|
255 (* Subtree versus frontier: *) |
|
256 lemma subtr_inFr: |
|
257 assumes "inFr ns tr t" and "subtr ns tr tr1" |
|
258 shows "inFr ns tr1 t" |
|
259 proof- |
|
260 have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inFr ns tr t \<longrightarrow> inFr ns tr1 t)" |
|
261 apply(induct rule: subtr.induct, safe) by (metis inFr.Ind) |
|
262 thus ?thesis using assms by auto |
|
263 qed |
|
264 |
|
265 corollary Fr_subtr: |
|
266 "Fr ns tr = \<Union> {Fr ns tr' | tr'. subtr ns tr' tr}" |
|
267 unfolding Fr_def proof safe |
|
268 fix t assume t: "inFr ns tr t" hence "root tr \<in> ns" by (rule inFr_root_in) |
|
269 thus "t \<in> \<Union>{{t. inFr ns tr' t} |tr'. subtr ns tr' tr}" |
|
270 apply(intro UnionI[of "{t. inFr ns tr t}" _ t]) using t subtr.Refl by auto |
|
271 qed(metis subtr_inFr) |
|
272 |
|
273 lemma inFr_subtr: |
|
274 assumes "inFr ns tr t" |
|
275 shows "\<exists> tr'. subtr ns tr' tr \<and> Inl t \<in> cont tr'" |
|
276 using assms apply(induct rule: inFr.induct) apply safe |
|
277 apply (metis subtr.Refl) |
|
278 by (metis (lifting) subtr.Step) |
|
279 |
|
280 corollary Fr_subtr_cont: |
|
281 "Fr ns tr = \<Union> {Inl -` cont tr' | tr'. subtr ns tr' tr}" |
|
282 unfolding Fr_def |
|
283 apply safe |
|
284 apply (frule inFr_subtr) |
|
285 apply auto |
|
286 by (metis inFr.Base subtr_inFr subtr_rootL_in) |
|
287 |
|
288 (* Subtree versus interior: *) |
|
289 lemma subtr_inItr: |
|
290 assumes "inItr ns tr n" and "subtr ns tr tr1" |
|
291 shows "inItr ns tr1 n" |
|
292 proof- |
|
293 have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inItr ns tr n \<longrightarrow> inItr ns tr1 n)" |
|
294 apply(induct rule: subtr.induct, safe) by (metis inItr.Ind) |
|
295 thus ?thesis using assms by auto |
|
296 qed |
|
297 |
|
298 corollary Itr_subtr: |
|
299 "Itr ns tr = \<Union> {Itr ns tr' | tr'. subtr ns tr' tr}" |
|
300 unfolding Itr_def apply safe |
|
301 apply (metis (lifting, mono_tags) UnionI inItr_root_in mem_Collect_eq subtr.Refl) |
|
302 by (metis subtr_inItr) |
|
303 |
|
304 lemma inItr_subtr: |
|
305 assumes "inItr ns tr n" |
|
306 shows "\<exists> tr'. subtr ns tr' tr \<and> root tr' = n" |
|
307 using assms apply(induct rule: inItr.induct) apply safe |
|
308 apply (metis subtr.Refl) |
|
309 by (metis (lifting) subtr.Step) |
|
310 |
|
311 corollary Itr_subtr_cont: |
|
312 "Itr ns tr = {root tr' | tr'. subtr ns tr' tr}" |
|
313 unfolding Itr_def apply safe |
|
314 apply (metis (lifting, mono_tags) UnionI inItr_subtr mem_Collect_eq vimageI2) |
|
315 by (metis inItr.Base subtr_inItr subtr_rootL_in) |
|
316 |
|
317 |
|
318 subsection{* The immediate subtree function *} |
|
319 |
|
320 (* production of: *) |
|
321 abbreviation "prodOf tr \<equiv> (id \<oplus> root) ` (cont tr)" |
|
322 (* subtree of: *) |
|
323 definition "subtrOf tr n \<equiv> SOME tr'. Inr tr' \<in> cont tr \<and> root tr' = n" |
|
324 |
|
325 lemma subtrOf: |
|
326 assumes n: "Inr n \<in> prodOf tr" |
|
327 shows "Inr (subtrOf tr n) \<in> cont tr \<and> root (subtrOf tr n) = n" |
|
328 proof- |
|
329 obtain tr' where "Inr tr' \<in> cont tr \<and> root tr' = n" |
|
330 using n unfolding image_def by (metis (lifting) Inr_oplus_elim assms) |
|
331 thus ?thesis unfolding subtrOf_def by(rule someI) |
|
332 qed |
|
333 |
|
334 lemmas Inr_subtrOf = subtrOf[THEN conjunct1] |
|
335 lemmas root_subtrOf[simp] = subtrOf[THEN conjunct2] |
|
336 |
|
337 lemma Inl_prodOf: "Inl -` (prodOf tr) = Inl -` (cont tr)" |
|
338 proof safe |
|
339 fix t ttr assume "Inl t = (id \<oplus> root) ttr" and "ttr \<in> cont tr" |
|
340 thus "t \<in> Inl -` cont tr" by(cases ttr, auto) |
|
341 next |
|
342 fix t assume "Inl t \<in> cont tr" thus "t \<in> Inl -` prodOf tr" |
|
343 by (metis (lifting) id_def image_iff sum_map.simps(1) vimageI2) |
|
344 qed |
|
345 |
|
346 lemma root_prodOf: |
|
347 assumes "Inr tr' \<in> cont tr" |
|
348 shows "Inr (root tr') \<in> prodOf tr" |
|
349 by (metis (lifting) assms image_iff sum_map.simps(2)) |
|
350 |
|
351 |
|
352 subsection{* Derivation trees *} |
|
353 |
|
354 coinductive dtree where |
|
355 Tree: "\<lbrakk>(root tr, (id \<oplus> root) ` (cont tr)) \<in> P; inj_on root (Inr -` cont tr); |
|
356 lift dtree (cont tr)\<rbrakk> \<Longrightarrow> dtree tr" |
|
357 monos lift_mono |
|
358 |
|
359 (* destruction rules: *) |
|
360 lemma dtree_P: |
|
361 assumes "dtree tr" |
|
362 shows "(root tr, (id \<oplus> root) ` (cont tr)) \<in> P" |
|
363 using assms unfolding dtree.simps by auto |
|
364 |
|
365 lemma dtree_inj_on: |
|
366 assumes "dtree tr" |
|
367 shows "inj_on root (Inr -` cont tr)" |
|
368 using assms unfolding dtree.simps by auto |
|
369 |
|
370 lemma dtree_inj[simp]: |
|
371 assumes "dtree tr" and "Inr tr1 \<in> cont tr" and "Inr tr2 \<in> cont tr" |
|
372 shows "root tr1 = root tr2 \<longleftrightarrow> tr1 = tr2" |
|
373 using assms dtree_inj_on unfolding inj_on_def by auto |
|
374 |
|
375 lemma dtree_lift: |
|
376 assumes "dtree tr" |
|
377 shows "lift dtree (cont tr)" |
|
378 using assms unfolding dtree.simps by auto |
|
379 |
|
380 |
|
381 (* coinduction:*) |
|
382 lemma dtree_coind[elim, consumes 1, case_names Hyp]: |
|
383 assumes phi: "\<phi> tr" |
|
384 and Hyp: |
|
385 "\<And> tr. \<phi> tr \<Longrightarrow> |
|
386 (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and> |
|
387 inj_on root (Inr -` cont tr) \<and> |
|
388 lift (\<lambda> tr. \<phi> tr \<or> dtree tr) (cont tr)" |
|
389 shows "dtree tr" |
|
390 apply(rule dtree.coinduct[of \<phi> tr, OF phi]) |
|
391 using Hyp by blast |
|
392 |
|
393 lemma dtree_raw_coind[elim, consumes 1, case_names Hyp]: |
|
394 assumes phi: "\<phi> tr" |
|
395 and Hyp: |
|
396 "\<And> tr. \<phi> tr \<Longrightarrow> |
|
397 (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and> |
|
398 inj_on root (Inr -` cont tr) \<and> |
|
399 lift \<phi> (cont tr)" |
|
400 shows "dtree tr" |
|
401 using phi apply(induct rule: dtree_coind) |
|
402 using Hyp mono_lift |
|
403 by (metis (mono_tags) mono_lift) |
|
404 |
|
405 lemma dtree_subtr_inj_on: |
|
406 assumes d: "dtree tr1" and s: "subtr ns tr tr1" |
|
407 shows "inj_on root (Inr -` cont tr)" |
|
408 using s d apply(induct rule: subtr.induct) |
|
409 apply (metis (lifting) dtree_inj_on) by (metis dtree_lift lift_def) |
|
410 |
|
411 lemma dtree_subtr_P: |
|
412 assumes d: "dtree tr1" and s: "subtr ns tr tr1" |
|
413 shows "(root tr, (id \<oplus> root) ` cont tr) \<in> P" |
|
414 using s d apply(induct rule: subtr.induct) |
|
415 apply (metis (lifting) dtree_P) by (metis dtree_lift lift_def) |
|
416 |
|
417 lemma subtrOf_root[simp]: |
|
418 assumes tr: "dtree tr" and cont: "Inr tr' \<in> cont tr" |
|
419 shows "subtrOf tr (root tr') = tr'" |
|
420 proof- |
|
421 have 0: "Inr (subtrOf tr (root tr')) \<in> cont tr" using Inr_subtrOf |
|
422 by (metis (lifting) cont root_prodOf) |
|
423 have "root (subtrOf tr (root tr')) = root tr'" |
|
424 using root_subtrOf by (metis (lifting) cont root_prodOf) |
|
425 thus ?thesis unfolding dtree_inj[OF tr 0 cont] . |
|
426 qed |
|
427 |
|
428 lemma surj_subtrOf: |
|
429 assumes "dtree tr" and 0: "Inr tr' \<in> cont tr" |
|
430 shows "\<exists> n. Inr n \<in> prodOf tr \<and> subtrOf tr n = tr'" |
|
431 apply(rule exI[of _ "root tr'"]) |
|
432 using root_prodOf[OF 0] subtrOf_root[OF assms] by simp |
|
433 |
|
434 lemma dtree_subtr: |
|
435 assumes "dtree tr1" and "subtr ns tr tr1" |
|
436 shows "dtree tr" |
|
437 proof- |
|
438 have "(\<exists> ns tr1. dtree tr1 \<and> subtr ns tr tr1) \<Longrightarrow> dtree tr" |
|
439 proof (induct rule: dtree_raw_coind) |
|
440 case (Hyp tr) |
|
441 then obtain ns tr1 where tr1: "dtree tr1" and tr_tr1: "subtr ns tr tr1" by auto |
|
442 show ?case unfolding lift_def proof safe |
|
443 show "(root tr, (id \<oplus> root) ` cont tr) \<in> P" using dtree_subtr_P[OF tr1 tr_tr1] . |
|
444 next |
|
445 show "inj_on root (Inr -` cont tr)" using dtree_subtr_inj_on[OF tr1 tr_tr1] . |
|
446 next |
|
447 fix tr' assume tr': "Inr tr' \<in> cont tr" |
|
448 have tr_tr1: "subtr (ns \<union> {root tr'}) tr tr1" using subtr_mono[OF tr_tr1] by auto |
|
449 have "subtr (ns \<union> {root tr'}) tr' tr1" using subtr_StepL[OF _ tr' tr_tr1] by auto |
|
450 thus "\<exists>ns' tr1. dtree tr1 \<and> subtr ns' tr' tr1" using tr1 by blast |
|
451 qed |
|
452 qed |
|
453 thus ?thesis using assms by auto |
|
454 qed |
|
455 |
|
456 |
|
457 subsection{* Default trees *} |
|
458 |
|
459 (* Pick a left-hand side of a production for each nonterminal *) |
|
460 definition S where "S n \<equiv> SOME tns. (n,tns) \<in> P" |
|
461 |
|
462 lemma S_P: "(n, S n) \<in> P" |
|
463 using used unfolding S_def by(rule someI_ex) |
|
464 |
|
465 lemma finite_S: "finite (S n)" |
|
466 using S_P finite_in_P by auto |
|
467 |
|
468 |
|
469 (* The default tree of a nonterminal *) |
|
470 definition deftr :: "N \<Rightarrow> Tree" where |
|
471 "deftr \<equiv> unfold id S" |
|
472 |
|
473 lemma deftr_simps[simp]: |
|
474 "root (deftr n) = n" |
|
475 "cont (deftr n) = image (id \<oplus> deftr) (S n)" |
|
476 using unfold(1)[of id S n] unfold(2)[of S n id, OF finite_S] |
|
477 unfolding deftr_def by simp_all |
|
478 |
|
479 lemmas root_deftr = deftr_simps(1) |
|
480 lemmas cont_deftr = deftr_simps(2) |
|
481 |
|
482 lemma root_o_deftr[simp]: "root o deftr = id" |
|
483 by (rule ext, auto) |
|
484 |
|
485 lemma dtree_deftr: "dtree (deftr n)" |
|
486 proof- |
|
487 {fix tr assume "\<exists> n. tr = deftr n" hence "dtree tr" |
|
488 apply(induct rule: dtree_raw_coind) apply safe |
|
489 unfolding deftr_simps image_compose[symmetric] sum_map.comp id_o |
|
490 root_o_deftr sum_map.id image_id id_apply apply(rule S_P) |
|
491 unfolding inj_on_def lift_def by auto |
|
492 } |
|
493 thus ?thesis by auto |
|
494 qed |
|
495 |
|
496 |
|
497 subsection{* Hereditary substitution *} |
|
498 |
|
499 (* Auxiliary concept: The root-ommiting frontier: *) |
|
500 definition "inFrr ns tr t \<equiv> \<exists> tr'. Inr tr' \<in> cont tr \<and> inFr ns tr' t" |
|
501 definition "Frr ns tr \<equiv> {t. \<exists> tr'. Inr tr' \<in> cont tr \<and> t \<in> Fr ns tr'}" |
|
502 |
|
503 context |
|
504 fixes tr0 :: Tree |
|
505 begin |
|
506 |
|
507 definition "hsubst_r tr \<equiv> root tr" |
|
508 definition "hsubst_c tr \<equiv> if root tr = root tr0 then cont tr0 else cont tr" |
|
509 |
|
510 (* Hereditary substitution: *) |
|
511 definition hsubst :: "Tree \<Rightarrow> Tree" where |
|
512 "hsubst \<equiv> unfold hsubst_r hsubst_c" |
|
513 |
|
514 lemma finite_hsubst_c: "finite (hsubst_c n)" |
|
515 unfolding hsubst_c_def by (metis finite_cont) |
|
516 |
|
517 lemma root_hsubst[simp]: "root (hsubst tr) = root tr" |
|
518 using unfold(1)[of hsubst_r hsubst_c tr] unfolding hsubst_def hsubst_r_def by simp |
|
519 |
|
520 lemma root_o_subst[simp]: "root o hsubst = root" |
|
521 unfolding comp_def root_hsubst .. |
|
522 |
|
523 lemma cont_hsubst_eq[simp]: |
|
524 assumes "root tr = root tr0" |
|
525 shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr0)" |
|
526 apply(subst id_o[symmetric, of id]) unfolding id_o |
|
527 using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c] |
|
528 unfolding hsubst_def hsubst_c_def using assms by simp |
|
529 |
|
530 lemma hsubst_eq: |
|
531 assumes "root tr = root tr0" |
|
532 shows "hsubst tr = hsubst tr0" |
|
533 apply(rule Tree_cong) using assms cont_hsubst_eq by auto |
|
534 |
|
535 lemma cont_hsubst_neq[simp]: |
|
536 assumes "root tr \<noteq> root tr0" |
|
537 shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr)" |
|
538 apply(subst id_o[symmetric, of id]) unfolding id_o |
|
539 using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c] |
|
540 unfolding hsubst_def hsubst_c_def using assms by simp |
|
541 |
|
542 lemma Inl_cont_hsubst_eq[simp]: |
|
543 assumes "root tr = root tr0" |
|
544 shows "Inl -` cont (hsubst tr) = Inl -` (cont tr0)" |
|
545 unfolding cont_hsubst_eq[OF assms] by simp |
|
546 |
|
547 lemma Inr_cont_hsubst_eq[simp]: |
|
548 assumes "root tr = root tr0" |
|
549 shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr0" |
|
550 unfolding cont_hsubst_eq[OF assms] by simp |
|
551 |
|
552 lemma Inl_cont_hsubst_neq[simp]: |
|
553 assumes "root tr \<noteq> root tr0" |
|
554 shows "Inl -` cont (hsubst tr) = Inl -` (cont tr)" |
|
555 unfolding cont_hsubst_neq[OF assms] by simp |
|
556 |
|
557 lemma Inr_cont_hsubst_neq[simp]: |
|
558 assumes "root tr \<noteq> root tr0" |
|
559 shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr" |
|
560 unfolding cont_hsubst_neq[OF assms] by simp |
|
561 |
|
562 lemma dtree_hsubst: |
|
563 assumes tr0: "dtree tr0" and tr: "dtree tr" |
|
564 shows "dtree (hsubst tr)" |
|
565 proof- |
|
566 {fix tr1 have "(\<exists> tr. dtree tr \<and> tr1 = hsubst tr) \<Longrightarrow> dtree tr1" |
|
567 proof (induct rule: dtree_raw_coind) |
|
568 case (Hyp tr1) then obtain tr |
|
569 where dtr: "dtree tr" and tr1: "tr1 = hsubst tr" by auto |
|
570 show ?case unfolding lift_def tr1 proof safe |
|
571 show "(root (hsubst tr), prodOf (hsubst tr)) \<in> P" |
|
572 unfolding tr1 apply(cases "root tr = root tr0") |
|
573 using dtree_P[OF dtr] dtree_P[OF tr0] |
|
574 by (auto simp add: image_compose[symmetric] sum_map.comp) |
|
575 show "inj_on root (Inr -` cont (hsubst tr))" |
|
576 apply(cases "root tr = root tr0") using dtree_inj_on[OF dtr] dtree_inj_on[OF tr0] |
|
577 unfolding inj_on_def by (auto, blast) |
|
578 fix tr' assume "Inr tr' \<in> cont (hsubst tr)" |
|
579 thus "\<exists>tra. dtree tra \<and> tr' = hsubst tra" |
|
580 apply(cases "root tr = root tr0", simp_all) |
|
581 apply (metis dtree_lift lift_def tr0) |
|
582 by (metis dtr dtree_lift lift_def) |
|
583 qed |
|
584 qed |
|
585 } |
|
586 thus ?thesis using assms by blast |
|
587 qed |
|
588 |
|
589 lemma Frr: "Frr ns tr = {t. inFrr ns tr t}" |
|
590 unfolding inFrr_def Frr_def Fr_def by auto |
|
591 |
|
592 lemma inFr_hsubst_imp: |
|
593 assumes "inFr ns (hsubst tr) t" |
|
594 shows "t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or> |
|
595 inFr (ns - {root tr0}) tr t" |
|
596 proof- |
|
597 {fix tr1 |
|
598 have "inFr ns tr1 t \<Longrightarrow> |
|
599 (\<And> tr. tr1 = hsubst tr \<Longrightarrow> (t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or> |
|
600 inFr (ns - {root tr0}) tr t))" |
|
601 proof(induct rule: inFr.induct) |
|
602 case (Base tr1 ns t tr) |
|
603 hence rtr: "root tr1 \<in> ns" and t_tr1: "Inl t \<in> cont tr1" and tr1: "tr1 = hsubst tr" |
|
604 by auto |
|
605 show ?case |
|
606 proof(cases "root tr1 = root tr0") |
|
607 case True |
|
608 hence "t \<in> Inl -` (cont tr0)" using t_tr1 unfolding tr1 by auto |
|
609 thus ?thesis by simp |
|
610 next |
|
611 case False |
|
612 hence "inFr (ns - {root tr0}) tr t" using t_tr1 unfolding tr1 apply simp |
|
613 by (metis Base.prems Diff_iff root_hsubst inFr.Base rtr singletonE) |
|
614 thus ?thesis by simp |
|
615 qed |
|
616 next |
|
617 case (Ind tr1 ns tr1' t) note IH = Ind(4) |
|
618 have rtr1: "root tr1 \<in> ns" and tr1'_tr1: "Inr tr1' \<in> cont tr1" |
|
619 and t_tr1': "inFr ns tr1' t" and tr1: "tr1 = hsubst tr" using Ind by auto |
|
620 have rtr1: "root tr1 = root tr" unfolding tr1 by simp |
|
621 show ?case |
|
622 proof(cases "root tr1 = root tr0") |
|
623 case True |
|
624 then obtain tr' where tr'_tr0: "Inr tr' \<in> cont tr0" and tr1': "tr1' = hsubst tr'" |
|
625 using tr1'_tr1 unfolding tr1 by auto |
|
626 show ?thesis using IH[OF tr1'] proof (elim disjE) |
|
627 assume "inFr (ns - {root tr0}) tr' t" |
|
628 thus ?thesis using tr'_tr0 unfolding inFrr_def by auto |
|
629 qed auto |
|
630 next |
|
631 case False |
|
632 then obtain tr' where tr'_tr: "Inr tr' \<in> cont tr" and tr1': "tr1' = hsubst tr'" |
|
633 using tr1'_tr1 unfolding tr1 by auto |
|
634 show ?thesis using IH[OF tr1'] proof (elim disjE) |
|
635 assume "inFr (ns - {root tr0}) tr' t" |
|
636 thus ?thesis using tr'_tr unfolding inFrr_def |
|
637 by (metis Diff_iff False Ind(1) empty_iff inFr2_Ind inFr_inFr2 insert_iff rtr1) |
|
638 qed auto |
|
639 qed |
|
640 qed |
|
641 } |
|
642 thus ?thesis using assms by auto |
|
643 qed |
|
644 |
|
645 lemma inFr_hsubst_notin: |
|
646 assumes "inFr ns tr t" and "root tr0 \<notin> ns" |
|
647 shows "inFr ns (hsubst tr) t" |
|
648 using assms apply(induct rule: inFr.induct) |
|
649 apply (metis Inl_cont_hsubst_neq inFr2.Base inFr_inFr2 root_hsubst vimageD vimageI2) |
|
650 by (metis (lifting) Inr_cont_hsubst_neq inFr.Ind rev_image_eqI root_hsubst vimageD vimageI2) |
|
651 |
|
652 lemma inFr_hsubst_minus: |
|
653 assumes "inFr (ns - {root tr0}) tr t" |
|
654 shows "inFr ns (hsubst tr) t" |
|
655 proof- |
|
656 have 1: "inFr (ns - {root tr0}) (hsubst tr) t" |
|
657 using inFr_hsubst_notin[OF assms] by simp |
|
658 show ?thesis using inFr_mono[OF 1] by auto |
|
659 qed |
|
660 |
|
661 lemma inFr_self_hsubst: |
|
662 assumes "root tr0 \<in> ns" |
|
663 shows |
|
664 "inFr ns (hsubst tr0) t \<longleftrightarrow> |
|
665 t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t" |
|
666 (is "?A \<longleftrightarrow> ?B \<or> ?C") |
|
667 apply(intro iffI) |
|
668 apply (metis inFr_hsubst_imp Diff_iff inFr_root_in insertI1) proof(elim disjE) |
|
669 assume ?B thus ?A apply(intro inFr.Base) using assms by auto |
|
670 next |
|
671 assume ?C then obtain tr where |
|
672 tr_tr0: "Inr tr \<in> cont tr0" and t_tr: "inFr (ns - {root tr0}) tr t" |
|
673 unfolding inFrr_def by auto |
|
674 def tr1 \<equiv> "hsubst tr" |
|
675 have 1: "inFr ns tr1 t" using t_tr unfolding tr1_def using inFr_hsubst_minus by auto |
|
676 have "Inr tr1 \<in> cont (hsubst tr0)" unfolding tr1_def using tr_tr0 by auto |
|
677 thus ?A using 1 inFr.Ind assms by (metis root_hsubst) |
|
678 qed |
|
679 |
|
680 theorem Fr_self_hsubst: |
|
681 assumes "root tr0 \<in> ns" |
|
682 shows "Fr ns (hsubst tr0) = Inl -` (cont tr0) \<union> Frr (ns - {root tr0}) tr0" |
|
683 using inFr_self_hsubst[OF assms] unfolding Frr Fr_def by auto |
|
684 |
|
685 end (* context *) |
|
686 |
|
687 |
|
688 subsection{* Regular trees *} |
|
689 |
|
690 hide_const regular |
|
691 |
|
692 definition "reg f tr \<equiv> \<forall> tr'. subtr UNIV tr' tr \<longrightarrow> tr' = f (root tr')" |
|
693 definition "regular tr \<equiv> \<exists> f. reg f tr" |
|
694 |
|
695 lemma reg_def2: "reg f tr \<longleftrightarrow> (\<forall> ns tr'. subtr ns tr' tr \<longrightarrow> tr' = f (root tr'))" |
|
696 unfolding reg_def using subtr_mono by (metis subset_UNIV) |
|
697 |
|
698 lemma regular_def2: "regular tr \<longleftrightarrow> (\<exists> f. reg f tr \<and> (\<forall> n. root (f n) = n))" |
|
699 unfolding regular_def proof safe |
|
700 fix f assume f: "reg f tr" |
|
701 def g \<equiv> "\<lambda> n. if inItr UNIV tr n then f n else deftr n" |
|
702 show "\<exists>g. reg g tr \<and> (\<forall>n. root (g n) = n)" |
|
703 apply(rule exI[of _ g]) |
|
704 using f deftr_simps(1) unfolding g_def reg_def apply safe |
|
705 apply (metis (lifting) inItr.Base subtr_inItr subtr_rootL_in) |
|
706 by (metis (full_types) inItr_subtr subtr_subtr2) |
|
707 qed auto |
|
708 |
|
709 lemma reg_root: |
|
710 assumes "reg f tr" |
|
711 shows "f (root tr) = tr" |
|
712 using assms unfolding reg_def |
|
713 by (metis (lifting) iso_tuple_UNIV_I subtr.Refl) |
|
714 |
|
715 |
|
716 lemma reg_Inr_cont: |
|
717 assumes "reg f tr" and "Inr tr' \<in> cont tr" |
|
718 shows "reg f tr'" |
|
719 by (metis (lifting) assms iso_tuple_UNIV_I reg_def subtr.Step) |
|
720 |
|
721 lemma reg_subtr: |
|
722 assumes "reg f tr" and "subtr ns tr' tr" |
|
723 shows "reg f tr'" |
|
724 using assms unfolding reg_def using subtr_trans[of UNIV tr] UNIV_I |
|
725 by (metis UNIV_eq_I UnCI Un_upper1 iso_tuple_UNIV_I subtr_mono subtr_trans) |
|
726 |
|
727 lemma regular_subtr: |
|
728 assumes r: "regular tr" and s: "subtr ns tr' tr" |
|
729 shows "regular tr'" |
|
730 using r reg_subtr[OF _ s] unfolding regular_def by auto |
|
731 |
|
732 lemma subtr_deftr: |
|
733 assumes "subtr ns tr' (deftr n)" |
|
734 shows "tr' = deftr (root tr')" |
|
735 proof- |
|
736 {fix tr have "subtr ns tr' tr \<Longrightarrow> (\<forall> n. tr = deftr n \<longrightarrow> tr' = deftr (root tr'))" |
|
737 apply (induct rule: subtr.induct) |
|
738 proof(metis (lifting) deftr_simps(1), safe) |
|
739 fix tr3 ns tr1 tr2 n |
|
740 assume 1: "root (deftr n) \<in> ns" and 2: "subtr ns tr1 tr2" |
|
741 and IH: "\<forall>n. tr2 = deftr n \<longrightarrow> tr1 = deftr (root tr1)" |
|
742 and 3: "Inr tr2 \<in> cont (deftr n)" |
|
743 have "tr2 \<in> deftr ` UNIV" |
|
744 using 3 unfolding deftr_simps image_def |
|
745 by (metis (lifting, full_types) 3 CollectI Inr_oplus_iff cont_deftr |
|
746 iso_tuple_UNIV_I) |
|
747 then obtain n where "tr2 = deftr n" by auto |
|
748 thus "tr1 = deftr (root tr1)" using IH by auto |
|
749 qed |
|
750 } |
|
751 thus ?thesis using assms by auto |
|
752 qed |
|
753 |
|
754 lemma reg_deftr: "reg deftr (deftr n)" |
|
755 unfolding reg_def using subtr_deftr by auto |
|
756 |
|
757 lemma dtree_subtrOf_Union: |
|
758 assumes "dtree tr" |
|
759 shows "\<Union>{K tr' |tr'. Inr tr' \<in> cont tr} = |
|
760 \<Union>{K (subtrOf tr n) |n. Inr n \<in> prodOf tr}" |
|
761 unfolding Union_eq Bex_def mem_Collect_eq proof safe |
|
762 fix x xa tr' |
|
763 assume x: "x \<in> K tr'" and tr'_tr: "Inr tr' \<in> cont tr" |
|
764 show "\<exists>X. (\<exists>n. X = K (subtrOf tr n) \<and> Inr n \<in> prodOf tr) \<and> x \<in> X" |
|
765 apply(rule exI[of _ "K (subtrOf tr (root tr'))"]) apply(intro conjI) |
|
766 apply(rule exI[of _ "root tr'"]) apply (metis (lifting) root_prodOf tr'_tr) |
|
767 by (metis (lifting) assms subtrOf_root tr'_tr x) |
|
768 next |
|
769 fix x X n ttr |
|
770 assume x: "x \<in> K (subtrOf tr n)" and n: "Inr n = (id \<oplus> root) ttr" and ttr: "ttr \<in> cont tr" |
|
771 show "\<exists>X. (\<exists>tr'. X = K tr' \<and> Inr tr' \<in> cont tr) \<and> x \<in> X" |
|
772 apply(rule exI[of _ "K (subtrOf tr n)"]) apply(intro conjI) |
|
773 apply(rule exI[of _ "subtrOf tr n"]) apply (metis imageI n subtrOf ttr) |
|
774 using x . |
|
775 qed |
|
776 |
|
777 |
|
778 |
|
779 |
|
780 subsection {* Paths in a regular tree *} |
|
781 |
|
782 inductive path :: "(N \<Rightarrow> Tree) \<Rightarrow> N list \<Rightarrow> bool" for f where |
|
783 Base: "path f [n]" |
|
784 | |
|
785 Ind: "\<lbrakk>path f (n1 # nl); Inr (f n1) \<in> cont (f n)\<rbrakk> |
|
786 \<Longrightarrow> path f (n # n1 # nl)" |
|
787 |
|
788 lemma path_NE: |
|
789 assumes "path f nl" |
|
790 shows "nl \<noteq> Nil" |
|
791 using assms apply(induct rule: path.induct) by auto |
|
792 |
|
793 lemma path_post: |
|
794 assumes f: "path f (n # nl)" and nl: "nl \<noteq> []" |
|
795 shows "path f nl" |
|
796 proof- |
|
797 obtain n1 nl1 where nl: "nl = n1 # nl1" using nl by (cases nl, auto) |
|
798 show ?thesis using assms unfolding nl using path.simps by (metis (lifting) list.inject) |
|
799 qed |
|
800 |
|
801 lemma path_post_concat: |
|
802 assumes "path f (nl1 @ nl2)" and "nl2 \<noteq> Nil" |
|
803 shows "path f nl2" |
|
804 using assms apply (induct nl1) |
|
805 apply (metis append_Nil) by (metis Nil_is_append_conv append_Cons path_post) |
|
806 |
|
807 lemma path_concat: |
|
808 assumes "path f nl1" and "path f ((last nl1) # nl2)" |
|
809 shows "path f (nl1 @ nl2)" |
|
810 using assms apply(induct rule: path.induct) apply simp |
|
811 by (metis append_Cons last.simps list.simps(3) path.Ind) |
|
812 |
|
813 lemma path_distinct: |
|
814 assumes "path f nl" |
|
815 shows "\<exists> nl'. path f nl' \<and> hd nl' = hd nl \<and> last nl' = last nl \<and> |
|
816 set nl' \<subseteq> set nl \<and> distinct nl'" |
|
817 using assms proof(induct rule: length_induct) |
|
818 case (1 nl) hence p_nl: "path f nl" by simp |
|
819 then obtain n nl1 where nl: "nl = n # nl1" by (metis list.exhaust path_NE) |
|
820 show ?case |
|
821 proof(cases nl1) |
|
822 case Nil |
|
823 show ?thesis apply(rule exI[of _ nl]) using path.Base unfolding nl Nil by simp |
|
824 next |
|
825 case (Cons n1 nl2) |
|
826 hence p1: "path f nl1" by (metis list.simps nl p_nl path_post) |
|
827 show ?thesis |
|
828 proof(cases "n \<in> set nl1") |
|
829 case False |
|
830 obtain nl1' where p1': "path f nl1'" and hd_nl1': "hd nl1' = hd nl1" and |
|
831 l_nl1': "last nl1' = last nl1" and d_nl1': "distinct nl1'" |
|
832 and s_nl1': "set nl1' \<subseteq> set nl1" |
|
833 using 1(1)[THEN allE[of _ nl1]] p1 unfolding nl by auto |
|
834 obtain nl2' where nl1': "nl1' = n1 # nl2'" using path_NE[OF p1'] hd_nl1' |
|
835 unfolding Cons by(cases nl1', auto) |
|
836 show ?thesis apply(intro exI[of _ "n # nl1'"]) unfolding nl proof safe |
|
837 show "path f (n # nl1')" unfolding nl1' |
|
838 apply(rule path.Ind, metis nl1' p1') |
|
839 by (metis (lifting) Cons list.inject nl p1 p_nl path.simps path_NE) |
|
840 qed(insert l_nl1' Cons nl1' s_nl1' d_nl1' False, auto) |
|
841 next |
|
842 case True |
|
843 then obtain nl11 nl12 where nl1: "nl1 = nl11 @ n # nl12" |
|
844 by (metis split_list) |
|
845 have p12: "path f (n # nl12)" |
|
846 apply(rule path_post_concat[of _ "n # nl11"]) using p_nl[unfolded nl nl1] by auto |
|
847 obtain nl12' where p1': "path f nl12'" and hd_nl12': "hd nl12' = n" and |
|
848 l_nl12': "last nl12' = last (n # nl12)" and d_nl12': "distinct nl12'" |
|
849 and s_nl12': "set nl12' \<subseteq> {n} \<union> set nl12" |
|
850 using 1(1)[THEN allE[of _ "n # nl12"]] p12 unfolding nl nl1 by auto |
|
851 thus ?thesis apply(intro exI[of _ nl12']) unfolding nl nl1 by auto |
|
852 qed |
|
853 qed |
|
854 qed |
|
855 |
|
856 lemma path_subtr: |
|
857 assumes f: "\<And> n. root (f n) = n" |
|
858 and p: "path f nl" |
|
859 shows "subtr (set nl) (f (last nl)) (f (hd nl))" |
|
860 using p proof (induct rule: path.induct) |
|
861 case (Ind n1 nl n) let ?ns1 = "insert n1 (set nl)" |
|
862 have "path f (n1 # nl)" |
|
863 and "subtr ?ns1 (f (last (n1 # nl))) (f n1)" |
|
864 and fn1: "Inr (f n1) \<in> cont (f n)" using Ind by simp_all |
|
865 hence fn1_flast: "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n1)" |
|
866 by (metis subset_insertI subtr_mono) |
|
867 have 1: "last (n # n1 # nl) = last (n1 # nl)" by auto |
|
868 have "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n)" |
|
869 using f subtr.Step[OF _ fn1_flast fn1] by auto |
|
870 thus ?case unfolding 1 by simp |
|
871 qed (metis f hd.simps last_ConsL last_in_set not_Cons_self2 subtr.Refl) |
|
872 |
|
873 lemma reg_subtr_path_aux: |
|
874 assumes f: "reg f tr" and n: "subtr ns tr1 tr" |
|
875 shows "\<exists> nl. path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns" |
|
876 using n f proof(induct rule: subtr.induct) |
|
877 case (Refl tr ns) |
|
878 thus ?case |
|
879 apply(intro exI[of _ "[root tr]"]) apply simp by (metis (lifting) path.Base reg_root) |
|
880 next |
|
881 case (Step tr ns tr2 tr1) |
|
882 hence rtr: "root tr \<in> ns" and tr1_tr: "Inr tr1 \<in> cont tr" |
|
883 and tr2_tr1: "subtr ns tr2 tr1" and tr: "reg f tr" by auto |
|
884 have tr1: "reg f tr1" using reg_subtr[OF tr] rtr tr1_tr |
|
885 by (metis (lifting) Step.prems iso_tuple_UNIV_I reg_def subtr.Step) |
|
886 obtain nl where nl: "path f nl" and f_nl: "f (hd nl) = tr1" |
|
887 and last_nl: "f (last nl) = tr2" and set: "set nl \<subseteq> ns" using Step(3)[OF tr1] by auto |
|
888 have 0: "path f (root tr # nl)" apply (subst path.simps) |
|
889 using f_nl nl reg_root tr tr1_tr by (metis hd.simps neq_Nil_conv) |
|
890 show ?case apply(rule exI[of _ "(root tr) # nl"]) |
|
891 using 0 reg_root tr last_nl nl path_NE rtr set by auto |
|
892 qed |
|
893 |
|
894 lemma reg_subtr_path: |
|
895 assumes f: "reg f tr" and n: "subtr ns tr1 tr" |
|
896 shows "\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns" |
|
897 using reg_subtr_path_aux[OF assms] path_distinct[of f] |
|
898 by (metis (lifting) order_trans) |
|
899 |
|
900 lemma subtr_iff_path: |
|
901 assumes r: "reg f tr" and f: "\<And> n. root (f n) = n" |
|
902 shows "subtr ns tr1 tr \<longleftrightarrow> |
|
903 (\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns)" |
|
904 proof safe |
|
905 fix nl assume p: "path f nl" and nl: "set nl \<subseteq> ns" |
|
906 have "subtr (set nl) (f (last nl)) (f (hd nl))" |
|
907 apply(rule path_subtr) using p f by simp_all |
|
908 thus "subtr ns (f (last nl)) (f (hd nl))" |
|
909 using subtr_mono nl by auto |
|
910 qed(insert reg_subtr_path[OF r], auto) |
|
911 |
|
912 lemma inFr_iff_path: |
|
913 assumes r: "reg f tr" and f: "\<And> n. root (f n) = n" |
|
914 shows |
|
915 "inFr ns tr t \<longleftrightarrow> |
|
916 (\<exists> nl tr1. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> |
|
917 set nl \<subseteq> ns \<and> Inl t \<in> cont tr1)" |
|
918 apply safe |
|
919 apply (metis (no_types) inFr_subtr r reg_subtr_path) |
|
920 by (metis f inFr.Base path_subtr subtr_inFr subtr_mono subtr_rootL_in) |
|
921 |
|
922 |
|
923 |
|
924 subsection{* The regular cut of a tree *} |
|
925 |
|
926 context fixes tr0 :: Tree |
|
927 begin |
|
928 |
|
929 (* Picking a subtree of a certain root: *) |
|
930 definition "pick n \<equiv> SOME tr. subtr UNIV tr tr0 \<and> root tr = n" |
|
931 |
|
932 lemma pick: |
|
933 assumes "inItr UNIV tr0 n" |
|
934 shows "subtr UNIV (pick n) tr0 \<and> root (pick n) = n" |
|
935 proof- |
|
936 have "\<exists> tr. subtr UNIV tr tr0 \<and> root tr = n" |
|
937 using assms by (metis (lifting) inItr_subtr) |
|
938 thus ?thesis unfolding pick_def by(rule someI_ex) |
|
939 qed |
|
940 |
|
941 lemmas subtr_pick = pick[THEN conjunct1] |
|
942 lemmas root_pick = pick[THEN conjunct2] |
|
943 |
|
944 lemma dtree_pick: |
|
945 assumes tr0: "dtree tr0" and n: "inItr UNIV tr0 n" |
|
946 shows "dtree (pick n)" |
|
947 using dtree_subtr[OF tr0 subtr_pick[OF n]] . |
|
948 |
|
949 definition "regOf_r n \<equiv> root (pick n)" |
|
950 definition "regOf_c n \<equiv> (id \<oplus> root) ` cont (pick n)" |
|
951 |
|
952 (* The regular tree of a function: *) |
|
953 definition regOf :: "N \<Rightarrow> Tree" where |
|
954 "regOf \<equiv> unfold regOf_r regOf_c" |
|
955 |
|
956 lemma finite_regOf_c: "finite (regOf_c n)" |
|
957 unfolding regOf_c_def by (metis finite_cont finite_imageI) |
|
958 |
|
959 lemma root_regOf_pick: "root (regOf n) = root (pick n)" |
|
960 using unfold(1)[of regOf_r regOf_c n] unfolding regOf_def regOf_r_def by simp |
|
961 |
|
962 lemma root_regOf[simp]: |
|
963 assumes "inItr UNIV tr0 n" |
|
964 shows "root (regOf n) = n" |
|
965 unfolding root_regOf_pick root_pick[OF assms] .. |
|
966 |
|
967 lemma cont_regOf[simp]: |
|
968 "cont (regOf n) = (id \<oplus> (regOf o root)) ` cont (pick n)" |
|
969 apply(subst id_o[symmetric, of id]) unfolding sum_map.comp[symmetric] |
|
970 unfolding image_compose unfolding regOf_c_def[symmetric] |
|
971 using unfold(2)[of regOf_c n regOf_r, OF finite_regOf_c] |
|
972 unfolding regOf_def .. |
|
973 |
|
974 lemma Inl_cont_regOf[simp]: |
|
975 "Inl -` (cont (regOf n)) = Inl -` (cont (pick n))" |
|
976 unfolding cont_regOf by simp |
|
977 |
|
978 lemma Inr_cont_regOf: |
|
979 "Inr -` (cont (regOf n)) = (regOf \<circ> root) ` (Inr -` cont (pick n))" |
|
980 unfolding cont_regOf by simp |
|
981 |
|
982 lemma subtr_regOf: |
|
983 assumes n: "inItr UNIV tr0 n" and "subtr UNIV tr1 (regOf n)" |
|
984 shows "\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = regOf n1" |
|
985 proof- |
|
986 {fix tr ns assume "subtr UNIV tr1 tr" |
|
987 hence "tr = regOf n \<longrightarrow> (\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = regOf n1)" |
|
988 proof (induct rule: subtr_UNIV_inductL) |
|
989 case (Step tr2 tr1 tr) |
|
990 show ?case proof |
|
991 assume "tr = regOf n" |
|
992 then obtain n1 where tr2: "Inr tr2 \<in> cont tr1" |
|
993 and tr1_tr: "subtr UNIV tr1 tr" and n1: "inItr UNIV tr0 n1" and tr1: "tr1 = regOf n1" |
|
994 using Step by auto |
|
995 obtain tr2' where tr2: "tr2 = regOf (root tr2')" |
|
996 and tr2': "Inr tr2' \<in> cont (pick n1)" |
|
997 using tr2 Inr_cont_regOf[of n1] |
|
998 unfolding tr1 image_def o_def using vimage_eq by auto |
|
999 have "inItr UNIV tr0 (root tr2')" |
|
1000 using inItr.Base inItr.Ind n1 pick subtr_inItr tr2' by (metis iso_tuple_UNIV_I) |
|
1001 thus "\<exists>n2. inItr UNIV tr0 n2 \<and> tr2 = regOf n2" using tr2 by blast |
|
1002 qed |
|
1003 qed(insert n, auto) |
|
1004 } |
|
1005 thus ?thesis using assms by auto |
|
1006 qed |
|
1007 |
|
1008 lemma root_regOf_root: |
|
1009 assumes n: "inItr UNIV tr0 n" and t_tr: "t_tr \<in> cont (pick n)" |
|
1010 shows "(id \<oplus> (root \<circ> regOf \<circ> root)) t_tr = (id \<oplus> root) t_tr" |
|
1011 using assms apply(cases t_tr) |
|
1012 apply (metis (lifting) sum_map.simps(1)) |
|
1013 using pick regOf_def regOf_r_def unfold(1) |
|
1014 inItr.Base o_apply subtr_StepL subtr_inItr sum_map.simps(2) |
|
1015 by (metis UNIV_I) |
|
1016 |
|
1017 lemma regOf_P: |
|
1018 assumes tr0: "dtree tr0" and n: "inItr UNIV tr0 n" |
|
1019 shows "(n, (id \<oplus> root) ` cont (regOf n)) \<in> P" (is "?L \<in> P") |
|
1020 proof- |
|
1021 have "?L = (n, (id \<oplus> root) ` cont (pick n))" |
|
1022 unfolding cont_regOf image_compose[symmetric] sum_map.comp id_o o_assoc |
|
1023 unfolding Pair_eq apply(rule conjI[OF refl]) apply(rule image_cong[OF refl]) |
|
1024 by(rule root_regOf_root[OF n]) |
|
1025 moreover have "... \<in> P" by (metis (lifting) dtree_pick root_pick dtree_P n tr0) |
|
1026 ultimately show ?thesis by simp |
|
1027 qed |
|
1028 |
|
1029 lemma dtree_regOf: |
|
1030 assumes tr0: "dtree tr0" and "inItr UNIV tr0 n" |
|
1031 shows "dtree (regOf n)" |
|
1032 proof- |
|
1033 {fix tr have "\<exists> n. inItr UNIV tr0 n \<and> tr = regOf n \<Longrightarrow> dtree tr" |
|
1034 proof (induct rule: dtree_raw_coind) |
|
1035 case (Hyp tr) |
|
1036 then obtain n where n: "inItr UNIV tr0 n" and tr: "tr = regOf n" by auto |
|
1037 show ?case unfolding lift_def apply safe |
|
1038 apply (metis (lifting) regOf_P root_regOf n tr tr0) |
|
1039 unfolding tr Inr_cont_regOf unfolding inj_on_def apply clarsimp using root_regOf |
|
1040 apply (metis UNIV_I inItr.Base n pick subtr2.simps subtr_inItr subtr_subtr2) |
|
1041 by (metis n subtr.Refl subtr_StepL subtr_regOf tr UNIV_I) |
|
1042 qed |
|
1043 } |
|
1044 thus ?thesis using assms by blast |
|
1045 qed |
|
1046 |
|
1047 (* The regular cut of a tree: *) |
|
1048 definition "rcut \<equiv> regOf (root tr0)" |
|
1049 |
|
1050 theorem reg_rcut: "reg regOf rcut" |
|
1051 unfolding reg_def rcut_def |
|
1052 by (metis inItr.Base root_regOf subtr_regOf UNIV_I) |
|
1053 |
|
1054 lemma rcut_reg: |
|
1055 assumes "reg regOf tr0" |
|
1056 shows "rcut = tr0" |
|
1057 using assms unfolding rcut_def reg_def by (metis subtr.Refl UNIV_I) |
|
1058 |
|
1059 theorem rcut_eq: "rcut = tr0 \<longleftrightarrow> reg regOf tr0" |
|
1060 using reg_rcut rcut_reg by metis |
|
1061 |
|
1062 theorem regular_rcut: "regular rcut" |
|
1063 using reg_rcut unfolding regular_def by blast |
|
1064 |
|
1065 theorem Fr_rcut: "Fr UNIV rcut \<subseteq> Fr UNIV tr0" |
|
1066 proof safe |
|
1067 fix t assume "t \<in> Fr UNIV rcut" |
|
1068 then obtain tr where t: "Inl t \<in> cont tr" and tr: "subtr UNIV tr (regOf (root tr0))" |
|
1069 using Fr_subtr[of UNIV "regOf (root tr0)"] unfolding rcut_def |
|
1070 by (metis (full_types) Fr_def inFr_subtr mem_Collect_eq) |
|
1071 obtain n where n: "inItr UNIV tr0 n" and tr: "tr = regOf n" using tr |
|
1072 by (metis (lifting) inItr.Base subtr_regOf UNIV_I) |
|
1073 have "Inl t \<in> cont (pick n)" using t using Inl_cont_regOf[of n] unfolding tr |
|
1074 by (metis (lifting) vimageD vimageI2) |
|
1075 moreover have "subtr UNIV (pick n) tr0" using subtr_pick[OF n] .. |
|
1076 ultimately show "t \<in> Fr UNIV tr0" unfolding Fr_subtr_cont by auto |
|
1077 qed |
|
1078 |
|
1079 theorem dtree_rcut: |
|
1080 assumes "dtree tr0" |
|
1081 shows "dtree rcut" |
|
1082 unfolding rcut_def using dtree_regOf[OF assms inItr.Base] by simp |
|
1083 |
|
1084 theorem root_rcut[simp]: "root rcut = root tr0" |
|
1085 unfolding rcut_def |
|
1086 by (metis (lifting) root_regOf inItr.Base reg_def reg_root subtr_rootR_in) |
|
1087 |
|
1088 end (* context *) |
|
1089 |
|
1090 |
|
1091 subsection{* Recursive description of the regular tree frontiers *} |
|
1092 |
|
1093 lemma regular_inFr: |
|
1094 assumes r: "regular tr" and In: "root tr \<in> ns" |
|
1095 and t: "inFr ns tr t" |
|
1096 shows "t \<in> Inl -` (cont tr) \<or> |
|
1097 (\<exists> tr'. Inr tr' \<in> cont tr \<and> inFr (ns - {root tr}) tr' t)" |
|
1098 (is "?L \<or> ?R") |
|
1099 proof- |
|
1100 obtain f where r: "reg f tr" and f: "\<And>n. root (f n) = n" |
|
1101 using r unfolding regular_def2 by auto |
|
1102 obtain nl tr1 where d_nl: "distinct nl" and p: "path f nl" and hd_nl: "f (hd nl) = tr" |
|
1103 and l_nl: "f (last nl) = tr1" and s_nl: "set nl \<subseteq> ns" and t_tr1: "Inl t \<in> cont tr1" |
|
1104 using t unfolding inFr_iff_path[OF r f] by auto |
|
1105 obtain n nl1 where nl: "nl = n # nl1" by (metis (lifting) p path.simps) |
|
1106 hence f_n: "f n = tr" using hd_nl by simp |
|
1107 have n_nl1: "n \<notin> set nl1" using d_nl unfolding nl by auto |
|
1108 show ?thesis |
|
1109 proof(cases nl1) |
|
1110 case Nil hence "tr = tr1" using f_n l_nl unfolding nl by simp |
|
1111 hence ?L using t_tr1 by simp thus ?thesis by simp |
|
1112 next |
|
1113 case (Cons n1 nl2) note nl1 = Cons |
|
1114 have 1: "last nl1 = last nl" "hd nl1 = n1" unfolding nl nl1 by simp_all |
|
1115 have p1: "path f nl1" and n1_tr: "Inr (f n1) \<in> cont tr" |
|
1116 using path.simps[of f nl] p f_n unfolding nl nl1 by auto |
|
1117 have r1: "reg f (f n1)" using reg_Inr_cont[OF r n1_tr] . |
|
1118 have 0: "inFr (set nl1) (f n1) t" unfolding inFr_iff_path[OF r1 f] |
|
1119 apply(intro exI[of _ nl1], intro exI[of _ tr1]) |
|
1120 using d_nl unfolding 1 l_nl unfolding nl using p1 t_tr1 by auto |
|
1121 have root_tr: "root tr = n" by (metis f f_n) |
|
1122 have "inFr (ns - {root tr}) (f n1) t" apply(rule inFr_mono[OF 0]) |
|
1123 using s_nl unfolding root_tr unfolding nl using n_nl1 by auto |
|
1124 thus ?thesis using n1_tr by auto |
|
1125 qed |
|
1126 qed |
|
1127 |
|
1128 theorem regular_Fr: |
|
1129 assumes r: "regular tr" and In: "root tr \<in> ns" |
|
1130 shows "Fr ns tr = |
|
1131 Inl -` (cont tr) \<union> |
|
1132 \<Union> {Fr (ns - {root tr}) tr' | tr'. Inr tr' \<in> cont tr}" |
|
1133 unfolding Fr_def |
|
1134 using In inFr.Base regular_inFr[OF assms] apply safe |
|
1135 apply (simp, metis (full_types) UnionI mem_Collect_eq) |
|
1136 apply simp |
|
1137 by (simp, metis (lifting) inFr_Ind_minus insert_Diff) |
|
1138 |
|
1139 |
|
1140 subsection{* The generated languages *} |
|
1141 |
|
1142 (* The (possibly inifinite tree) generated language *) |
|
1143 definition "L ns n \<equiv> {Fr ns tr | tr. dtree tr \<and> root tr = n}" |
|
1144 |
|
1145 (* The regular-tree generated language *) |
|
1146 definition "Lr ns n \<equiv> {Fr ns tr | tr. dtree tr \<and> root tr = n \<and> regular tr}" |
|
1147 |
|
1148 theorem L_rec_notin: |
|
1149 assumes "n \<notin> ns" |
|
1150 shows "L ns n = {{}}" |
|
1151 using assms unfolding L_def apply safe |
|
1152 using not_root_Fr apply force |
|
1153 apply(rule exI[of _ "deftr n"]) |
|
1154 by (metis (no_types) dtree_deftr not_root_Fr root_deftr) |
|
1155 |
|
1156 theorem Lr_rec_notin: |
|
1157 assumes "n \<notin> ns" |
|
1158 shows "Lr ns n = {{}}" |
|
1159 using assms unfolding Lr_def apply safe |
|
1160 using not_root_Fr apply force |
|
1161 apply(rule exI[of _ "deftr n"]) |
|
1162 by (metis (no_types) regular_def dtree_deftr not_root_Fr reg_deftr root_deftr) |
|
1163 |
|
1164 lemma dtree_subtrOf: |
|
1165 assumes "dtree tr" and "Inr n \<in> prodOf tr" |
|
1166 shows "dtree (subtrOf tr n)" |
|
1167 by (metis assms dtree_lift lift_def subtrOf) |
|
1168 |
|
1169 theorem Lr_rec_in: |
|
1170 assumes n: "n \<in> ns" |
|
1171 shows "Lr ns n \<subseteq> |
|
1172 {Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K. |
|
1173 (n,tns) \<in> P \<and> |
|
1174 (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n')}" |
|
1175 (is "Lr ns n \<subseteq> {?F tns K | tns K. (n,tns) \<in> P \<and> ?\<phi> tns K}") |
|
1176 proof safe |
|
1177 fix ts assume "ts \<in> Lr ns n" |
|
1178 then obtain tr where dtr: "dtree tr" and r: "root tr = n" and tr: "regular tr" |
|
1179 and ts: "ts = Fr ns tr" unfolding Lr_def by auto |
|
1180 def tns \<equiv> "(id \<oplus> root) ` (cont tr)" |
|
1181 def K \<equiv> "\<lambda> n'. Fr (ns - {n}) (subtrOf tr n')" |
|
1182 show "\<exists>tns K. ts = ?F tns K \<and> (n, tns) \<in> P \<and> ?\<phi> tns K" |
|
1183 apply(rule exI[of _ tns], rule exI[of _ K]) proof(intro conjI allI impI) |
|
1184 show "ts = Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns}" |
|
1185 unfolding ts regular_Fr[OF tr n[unfolded r[symmetric]]] |
|
1186 unfolding tns_def K_def r[symmetric] |
|
1187 unfolding Inl_prodOf dtree_subtrOf_Union[OF dtr] .. |
|
1188 show "(n, tns) \<in> P" unfolding tns_def r[symmetric] using dtree_P[OF dtr] . |
|
1189 fix n' assume "Inr n' \<in> tns" thus "K n' \<in> Lr (ns - {n}) n'" |
|
1190 unfolding K_def Lr_def mem_Collect_eq apply(intro exI[of _ "subtrOf tr n'"]) |
|
1191 using dtr tr apply(intro conjI refl) unfolding tns_def |
|
1192 apply(erule dtree_subtrOf[OF dtr]) |
|
1193 apply (metis subtrOf) |
|
1194 by (metis Inr_subtrOf UNIV_I regular_subtr subtr.simps) |
|
1195 qed |
|
1196 qed |
|
1197 |
|
1198 lemma hsubst_aux: |
|
1199 fixes n ftr tns |
|
1200 assumes n: "n \<in> ns" and tns: "finite tns" and |
|
1201 1: "\<And> n'. Inr n' \<in> tns \<Longrightarrow> dtree (ftr n')" |
|
1202 defines "tr \<equiv> Node n ((id \<oplus> ftr) ` tns)" defines "tr' \<equiv> hsubst tr tr" |
|
1203 shows "Fr ns tr' = Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}" |
|
1204 (is "_ = ?B") proof- |
|
1205 have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns" |
|
1206 unfolding tr_def using tns by auto |
|
1207 have Frr: "Frr (ns - {n}) tr = \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}" |
|
1208 unfolding Frr_def ctr by auto |
|
1209 have "Fr ns tr' = Inl -` (cont tr) \<union> Frr (ns - {n}) tr" |
|
1210 using Fr_self_hsubst[OF n[unfolded rtr[symmetric]]] unfolding tr'_def rtr .. |
|
1211 also have "... = ?B" unfolding ctr Frr by simp |
|
1212 finally show ?thesis . |
|
1213 qed |
|
1214 |
|
1215 theorem L_rec_in: |
|
1216 assumes n: "n \<in> ns" |
|
1217 shows " |
|
1218 {Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K. |
|
1219 (n,tns) \<in> P \<and> |
|
1220 (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n')} |
|
1221 \<subseteq> L ns n" |
|
1222 proof safe |
|
1223 fix tns K |
|
1224 assume P: "(n, tns) \<in> P" and 0: "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n'" |
|
1225 {fix n' assume "Inr n' \<in> tns" |
|
1226 hence "K n' \<in> L (ns - {n}) n'" using 0 by auto |
|
1227 hence "\<exists> tr'. K n' = Fr (ns - {n}) tr' \<and> dtree tr' \<and> root tr' = n'" |
|
1228 unfolding L_def mem_Collect_eq by auto |
|
1229 } |
|
1230 then obtain ftr where 0: "\<And> n'. Inr n' \<in> tns \<Longrightarrow> |
|
1231 K n' = Fr (ns - {n}) (ftr n') \<and> dtree (ftr n') \<and> root (ftr n') = n'" |
|
1232 by metis |
|
1233 def tr \<equiv> "Node n ((id \<oplus> ftr) ` tns)" def tr' \<equiv> "hsubst tr tr" |
|
1234 have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns" |
|
1235 unfolding tr_def by (simp, metis P cont_Node finite_imageI finite_in_P) |
|
1236 have prtr: "prodOf tr = tns" apply(rule Inl_Inr_image_cong) |
|
1237 unfolding ctr apply simp apply simp apply safe |
|
1238 using 0 unfolding image_def apply force apply simp by (metis 0 vimageI2) |
|
1239 have 1: "{K n' |n'. Inr n' \<in> tns} = {Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}" |
|
1240 using 0 by auto |
|
1241 have dtr: "dtree tr" apply(rule dtree.Tree) |
|
1242 apply (metis (lifting) P prtr rtr) |
|
1243 unfolding inj_on_def ctr lift_def using 0 by auto |
|
1244 hence dtr': "dtree tr'" unfolding tr'_def by (metis dtree_hsubst) |
|
1245 have tns: "finite tns" using finite_in_P P by simp |
|
1246 have "Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns} \<in> L ns n" |
|
1247 unfolding L_def mem_Collect_eq apply(intro exI[of _ tr'] conjI) |
|
1248 using dtr' 0 hsubst_aux[OF assms tns, of ftr] unfolding tr_def tr'_def by auto |
|
1249 thus "Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} \<in> L ns n" unfolding 1 . |
|
1250 qed |
|
1251 |
|
1252 lemma card_N: "(n::N) \<in> ns \<Longrightarrow> card (ns - {n}) < card ns" |
|
1253 by (metis finite_N Diff_UNIV Diff_infinite_finite card_Diff1_less finite.emptyI) |
|
1254 |
|
1255 function LL where |
|
1256 "LL ns n = |
|
1257 (if n \<notin> ns then {{}} else |
|
1258 {Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K. |
|
1259 (n,tns) \<in> P \<and> |
|
1260 (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n')})" |
|
1261 by(pat_completeness, auto) |
|
1262 termination apply(relation "inv_image (measure card) fst") |
|
1263 using card_N by auto |
|
1264 |
|
1265 declare LL.simps[code] (* TODO: Does code generation for LL work? *) |
|
1266 declare LL.simps[simp del] |
|
1267 |
|
1268 theorem Lr_LL: "Lr ns n \<subseteq> LL ns n" |
|
1269 proof (induct ns arbitrary: n rule: measure_induct[of card]) |
|
1270 case (1 ns n) show ?case proof(cases "n \<in> ns") |
|
1271 case False thus ?thesis unfolding Lr_rec_notin[OF False] by (simp add: LL.simps) |
|
1272 next |
|
1273 case True show ?thesis apply(rule subset_trans) |
|
1274 using Lr_rec_in[OF True] apply assumption |
|
1275 unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp |
|
1276 fix tns K |
|
1277 assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast |
|
1278 assume "(n, tns) \<in> P" |
|
1279 and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n'" |
|
1280 thus "\<exists>tnsa Ka. |
|
1281 Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} = |
|
1282 Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and> |
|
1283 (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> LL (ns - {n}) n')" |
|
1284 apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto |
|
1285 qed |
|
1286 qed |
|
1287 qed |
|
1288 |
|
1289 theorem LL_L: "LL ns n \<subseteq> L ns n" |
|
1290 proof (induct ns arbitrary: n rule: measure_induct[of card]) |
|
1291 case (1 ns n) show ?case proof(cases "n \<in> ns") |
|
1292 case False thus ?thesis unfolding L_rec_notin[OF False] by (simp add: LL.simps) |
|
1293 next |
|
1294 case True show ?thesis apply(rule subset_trans) |
|
1295 prefer 2 using L_rec_in[OF True] apply assumption |
|
1296 unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp |
|
1297 fix tns K |
|
1298 assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast |
|
1299 assume "(n, tns) \<in> P" |
|
1300 and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n'" |
|
1301 thus "\<exists>tnsa Ka. |
|
1302 Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} = |
|
1303 Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and> |
|
1304 (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> L (ns - {n}) n')" |
|
1305 apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto |
|
1306 qed |
|
1307 qed |
|
1308 qed |
|
1309 |
|
1310 (* The subsumpsion relation between languages *) |
|
1311 definition "subs L1 L2 \<equiv> \<forall> ts2 \<in> L2. \<exists> ts1 \<in> L1. ts1 \<subseteq> ts2" |
|
1312 |
|
1313 lemma incl_subs[simp]: "L2 \<subseteq> L1 \<Longrightarrow> subs L1 L2" |
|
1314 unfolding subs_def by auto |
|
1315 |
|
1316 lemma subs_refl[simp]: "subs L1 L1" unfolding subs_def by auto |
|
1317 |
|
1318 lemma subs_trans: "\<lbrakk>subs L1 L2; subs L2 L3\<rbrakk> \<Longrightarrow> subs L1 L3" |
|
1319 unfolding subs_def by (metis subset_trans) |
|
1320 |
|
1321 (* Language equivalence *) |
|
1322 definition "leqv L1 L2 \<equiv> subs L1 L2 \<and> subs L2 L1" |
|
1323 |
|
1324 lemma subs_leqv[simp]: "leqv L1 L2 \<Longrightarrow> subs L1 L2" |
|
1325 unfolding leqv_def by auto |
|
1326 |
|
1327 lemma subs_leqv_sym[simp]: "leqv L1 L2 \<Longrightarrow> subs L2 L1" |
|
1328 unfolding leqv_def by auto |
|
1329 |
|
1330 lemma leqv_refl[simp]: "leqv L1 L1" unfolding leqv_def by auto |
|
1331 |
|
1332 lemma leqv_trans: |
|
1333 assumes 12: "leqv L1 L2" and 23: "leqv L2 L3" |
|
1334 shows "leqv L1 L3" |
|
1335 using assms unfolding leqv_def by (metis (lifting) subs_trans) |
|
1336 |
|
1337 lemma leqv_sym: "leqv L1 L2 \<Longrightarrow> leqv L2 L1" |
|
1338 unfolding leqv_def by auto |
|
1339 |
|
1340 lemma leqv_Sym: "leqv L1 L2 \<longleftrightarrow> leqv L2 L1" |
|
1341 unfolding leqv_def by auto |
|
1342 |
|
1343 lemma Lr_incl_L: "Lr ns ts \<subseteq> L ns ts" |
|
1344 unfolding Lr_def L_def by auto |
|
1345 |
|
1346 lemma Lr_subs_L: "subs (Lr UNIV ts) (L UNIV ts)" |
|
1347 unfolding subs_def proof safe |
|
1348 fix ts2 assume "ts2 \<in> L UNIV ts" |
|
1349 then obtain tr where ts2: "ts2 = Fr UNIV tr" and dtr: "dtree tr" and rtr: "root tr = ts" |
|
1350 unfolding L_def by auto |
|
1351 thus "\<exists>ts1\<in>Lr UNIV ts. ts1 \<subseteq> ts2" |
|
1352 apply(intro bexI[of _ "Fr UNIV (rcut tr)"]) |
|
1353 unfolding Lr_def L_def using Fr_rcut dtree_rcut root_rcut regular_rcut by auto |
|
1354 qed |
|
1355 |
|
1356 theorem Lr_leqv_L: "leqv (Lr UNIV ts) (L UNIV ts)" |
|
1357 using Lr_subs_L unfolding leqv_def by (metis (lifting) Lr_incl_L incl_subs) |
|
1358 |
|
1359 theorem LL_leqv_L: "leqv (LL UNIV ts) (L UNIV ts)" |
|
1360 by (metis (lifting) LL_L Lr_LL Lr_subs_L incl_subs leqv_def subs_trans) |
|
1361 |
|
1362 theorem LL_leqv_Lr: "leqv (LL UNIV ts) (Lr UNIV ts)" |
|
1363 using Lr_leqv_L LL_leqv_L by (metis leqv_Sym leqv_trans) |
|
1364 |
|
1365 |
|
1366 end |