|
1 |
|
2 theory Positions |
|
3 imports PosixSpec Lexer |
|
4 begin |
|
5 |
|
6 chapter \<open>An alternative definition for POSIX values\<close> |
|
7 |
|
8 section \<open>Positions in Values\<close> |
|
9 |
|
10 fun |
|
11 at :: "val \<Rightarrow> nat list \<Rightarrow> val" |
|
12 where |
|
13 "at v [] = v" |
|
14 | "at (Left v) (0#ps)= at v ps" |
|
15 | "at (Right v) (Suc 0#ps)= at v ps" |
|
16 | "at (Seq v1 v2) (0#ps)= at v1 ps" |
|
17 | "at (Seq v1 v2) (Suc 0#ps)= at v2 ps" |
|
18 | "at (Stars vs) (n#ps)= at (nth vs n) ps" |
|
19 |
|
20 |
|
21 |
|
22 fun Pos :: "val \<Rightarrow> (nat list) set" |
|
23 where |
|
24 "Pos (Void) = {[]}" |
|
25 | "Pos (Char c) = {[]}" |
|
26 | "Pos (Left v) = {[]} \<union> {0#ps | ps. ps \<in> Pos v}" |
|
27 | "Pos (Right v) = {[]} \<union> {1#ps | ps. ps \<in> Pos v}" |
|
28 | "Pos (Seq v1 v2) = {[]} \<union> {0#ps | ps. ps \<in> Pos v1} \<union> {1#ps | ps. ps \<in> Pos v2}" |
|
29 | "Pos (Stars []) = {[]}" |
|
30 | "Pos (Stars (v#vs)) = {[]} \<union> {0#ps | ps. ps \<in> Pos v} \<union> {Suc n#ps | n ps. n#ps \<in> Pos (Stars vs)}" |
|
31 |
|
32 |
|
33 lemma Pos_stars: |
|
34 "Pos (Stars vs) = {[]} \<union> (\<Union>n < length vs. {n#ps | ps. ps \<in> Pos (vs ! n)})" |
|
35 apply(induct vs) |
|
36 apply(auto simp add: insert_ident less_Suc_eq_0_disj) |
|
37 done |
|
38 |
|
39 lemma Pos_empty: |
|
40 shows "[] \<in> Pos v" |
|
41 by (induct v rule: Pos.induct)(auto) |
|
42 |
|
43 |
|
44 abbreviation |
|
45 "intlen vs \<equiv> int (length vs)" |
|
46 |
|
47 |
|
48 definition pflat_len :: "val \<Rightarrow> nat list => int" |
|
49 where |
|
50 "pflat_len v p \<equiv> (if p \<in> Pos v then intlen (flat (at v p)) else -1)" |
|
51 |
|
52 lemma pflat_len_simps: |
|
53 shows "pflat_len (Seq v1 v2) (0#p) = pflat_len v1 p" |
|
54 and "pflat_len (Seq v1 v2) (Suc 0#p) = pflat_len v2 p" |
|
55 and "pflat_len (Left v) (0#p) = pflat_len v p" |
|
56 and "pflat_len (Left v) (Suc 0#p) = -1" |
|
57 and "pflat_len (Right v) (Suc 0#p) = pflat_len v p" |
|
58 and "pflat_len (Right v) (0#p) = -1" |
|
59 and "pflat_len (Stars (v#vs)) (Suc n#p) = pflat_len (Stars vs) (n#p)" |
|
60 and "pflat_len (Stars (v#vs)) (0#p) = pflat_len v p" |
|
61 and "pflat_len v [] = intlen (flat v)" |
|
62 by (auto simp add: pflat_len_def Pos_empty) |
|
63 |
|
64 lemma pflat_len_Stars_simps: |
|
65 assumes "n < length vs" |
|
66 shows "pflat_len (Stars vs) (n#p) = pflat_len (vs!n) p" |
|
67 using assms |
|
68 apply(induct vs arbitrary: n p) |
|
69 apply(auto simp add: less_Suc_eq_0_disj pflat_len_simps) |
|
70 done |
|
71 |
|
72 lemma pflat_len_outside: |
|
73 assumes "p \<notin> Pos v1" |
|
74 shows "pflat_len v1 p = -1 " |
|
75 using assms by (simp add: pflat_len_def) |
|
76 |
|
77 |
|
78 |
|
79 section \<open>Orderings\<close> |
|
80 |
|
81 |
|
82 definition prefix_list:: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<sqsubseteq>pre _" [60,59] 60) |
|
83 where |
|
84 "ps1 \<sqsubseteq>pre ps2 \<equiv> \<exists>ps'. ps1 @ps' = ps2" |
|
85 |
|
86 definition sprefix_list:: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<sqsubset>spre _" [60,59] 60) |
|
87 where |
|
88 "ps1 \<sqsubset>spre ps2 \<equiv> ps1 \<sqsubseteq>pre ps2 \<and> ps1 \<noteq> ps2" |
|
89 |
|
90 inductive lex_list :: "nat list \<Rightarrow> nat list \<Rightarrow> bool" ("_ \<sqsubset>lex _" [60,59] 60) |
|
91 where |
|
92 "[] \<sqsubset>lex (p#ps)" |
|
93 | "ps1 \<sqsubset>lex ps2 \<Longrightarrow> (p#ps1) \<sqsubset>lex (p#ps2)" |
|
94 | "p1 < p2 \<Longrightarrow> (p1#ps1) \<sqsubset>lex (p2#ps2)" |
|
95 |
|
96 lemma lex_irrfl: |
|
97 fixes ps1 ps2 :: "nat list" |
|
98 assumes "ps1 \<sqsubset>lex ps2" |
|
99 shows "ps1 \<noteq> ps2" |
|
100 using assms |
|
101 by(induct rule: lex_list.induct)(auto) |
|
102 |
|
103 lemma lex_simps [simp]: |
|
104 fixes xs ys :: "nat list" |
|
105 shows "[] \<sqsubset>lex ys \<longleftrightarrow> ys \<noteq> []" |
|
106 and "xs \<sqsubset>lex [] \<longleftrightarrow> False" |
|
107 and "(x # xs) \<sqsubset>lex (y # ys) \<longleftrightarrow> (x < y \<or> (x = y \<and> xs \<sqsubset>lex ys))" |
|
108 by (auto simp add: neq_Nil_conv elim: lex_list.cases intro: lex_list.intros) |
|
109 |
|
110 lemma lex_trans: |
|
111 fixes ps1 ps2 ps3 :: "nat list" |
|
112 assumes "ps1 \<sqsubset>lex ps2" "ps2 \<sqsubset>lex ps3" |
|
113 shows "ps1 \<sqsubset>lex ps3" |
|
114 using assms |
|
115 by (induct arbitrary: ps3 rule: lex_list.induct) |
|
116 (auto elim: lex_list.cases) |
|
117 |
|
118 |
|
119 lemma lex_trichotomous: |
|
120 fixes p q :: "nat list" |
|
121 shows "p = q \<or> p \<sqsubset>lex q \<or> q \<sqsubset>lex p" |
|
122 apply(induct p arbitrary: q) |
|
123 apply(auto elim: lex_list.cases) |
|
124 apply(case_tac q) |
|
125 apply(auto) |
|
126 done |
|
127 |
|
128 |
|
129 |
|
130 |
|
131 section \<open>POSIX Ordering of Values According to Okui \& Suzuki\<close> |
|
132 |
|
133 |
|
134 definition PosOrd:: "val \<Rightarrow> nat list \<Rightarrow> val \<Rightarrow> bool" ("_ \<sqsubset>val _ _" [60, 60, 59] 60) |
|
135 where |
|
136 "v1 \<sqsubset>val p v2 \<equiv> pflat_len v1 p > pflat_len v2 p \<and> |
|
137 (\<forall>q \<in> Pos v1 \<union> Pos v2. q \<sqsubset>lex p \<longrightarrow> pflat_len v1 q = pflat_len v2 q)" |
|
138 |
|
139 lemma PosOrd_def2: |
|
140 shows "v1 \<sqsubset>val p v2 \<longleftrightarrow> |
|
141 pflat_len v1 p > pflat_len v2 p \<and> |
|
142 (\<forall>q \<in> Pos v1. q \<sqsubset>lex p \<longrightarrow> pflat_len v1 q = pflat_len v2 q) \<and> |
|
143 (\<forall>q \<in> Pos v2. q \<sqsubset>lex p \<longrightarrow> pflat_len v1 q = pflat_len v2 q)" |
|
144 unfolding PosOrd_def |
|
145 apply(auto) |
|
146 done |
|
147 |
|
148 |
|
149 definition PosOrd_ex:: "val \<Rightarrow> val \<Rightarrow> bool" ("_ :\<sqsubset>val _" [60, 59] 60) |
|
150 where |
|
151 "v1 :\<sqsubset>val v2 \<equiv> \<exists>p. v1 \<sqsubset>val p v2" |
|
152 |
|
153 definition PosOrd_ex_eq:: "val \<Rightarrow> val \<Rightarrow> bool" ("_ :\<sqsubseteq>val _" [60, 59] 60) |
|
154 where |
|
155 "v1 :\<sqsubseteq>val v2 \<equiv> v1 :\<sqsubset>val v2 \<or> v1 = v2" |
|
156 |
|
157 |
|
158 lemma PosOrd_trans: |
|
159 assumes "v1 :\<sqsubset>val v2" "v2 :\<sqsubset>val v3" |
|
160 shows "v1 :\<sqsubset>val v3" |
|
161 proof - |
|
162 from assms obtain p p' |
|
163 where as: "v1 \<sqsubset>val p v2" "v2 \<sqsubset>val p' v3" unfolding PosOrd_ex_def by blast |
|
164 then have pos: "p \<in> Pos v1" "p' \<in> Pos v2" unfolding PosOrd_def pflat_len_def |
|
165 by (smt not_int_zless_negative)+ |
|
166 have "p = p' \<or> p \<sqsubset>lex p' \<or> p' \<sqsubset>lex p" |
|
167 by (rule lex_trichotomous) |
|
168 moreover |
|
169 { assume "p = p'" |
|
170 with as have "v1 \<sqsubset>val p v3" unfolding PosOrd_def pflat_len_def |
|
171 by (smt Un_iff) |
|
172 then have " v1 :\<sqsubset>val v3" unfolding PosOrd_ex_def by blast |
|
173 } |
|
174 moreover |
|
175 { assume "p \<sqsubset>lex p'" |
|
176 with as have "v1 \<sqsubset>val p v3" unfolding PosOrd_def pflat_len_def |
|
177 by (smt Un_iff lex_trans) |
|
178 then have " v1 :\<sqsubset>val v3" unfolding PosOrd_ex_def by blast |
|
179 } |
|
180 moreover |
|
181 { assume "p' \<sqsubset>lex p" |
|
182 with as have "v1 \<sqsubset>val p' v3" unfolding PosOrd_def |
|
183 by (smt Un_iff lex_trans pflat_len_def) |
|
184 then have "v1 :\<sqsubset>val v3" unfolding PosOrd_ex_def by blast |
|
185 } |
|
186 ultimately show "v1 :\<sqsubset>val v3" by blast |
|
187 qed |
|
188 |
|
189 lemma PosOrd_irrefl: |
|
190 assumes "v :\<sqsubset>val v" |
|
191 shows "False" |
|
192 using assms unfolding PosOrd_ex_def PosOrd_def |
|
193 by auto |
|
194 |
|
195 lemma PosOrd_assym: |
|
196 assumes "v1 :\<sqsubset>val v2" |
|
197 shows "\<not>(v2 :\<sqsubset>val v1)" |
|
198 using assms |
|
199 using PosOrd_irrefl PosOrd_trans by blast |
|
200 |
|
201 (* |
|
202 :\<sqsubseteq>val and :\<sqsubset>val are partial orders. |
|
203 *) |
|
204 |
|
205 lemma PosOrd_ordering: |
|
206 shows "ordering (\<lambda>v1 v2. v1 :\<sqsubseteq>val v2) (\<lambda> v1 v2. v1 :\<sqsubset>val v2)" |
|
207 unfolding ordering_def PosOrd_ex_eq_def |
|
208 apply(auto) |
|
209 using PosOrd_trans partial_preordering_def apply blast |
|
210 using PosOrd_assym ordering_axioms_def by blast |
|
211 |
|
212 lemma PosOrd_order: |
|
213 shows "class.order (\<lambda>v1 v2. v1 :\<sqsubseteq>val v2) (\<lambda> v1 v2. v1 :\<sqsubset>val v2)" |
|
214 using PosOrd_ordering |
|
215 apply(simp add: class.order_def class.preorder_def class.order_axioms_def) |
|
216 by (metis (full_types) PosOrd_ex_eq_def PosOrd_irrefl PosOrd_trans) |
|
217 |
|
218 |
|
219 lemma PosOrd_ex_eq2: |
|
220 shows "v1 :\<sqsubset>val v2 \<longleftrightarrow> (v1 :\<sqsubseteq>val v2 \<and> v1 \<noteq> v2)" |
|
221 using PosOrd_ordering |
|
222 using PosOrd_ex_eq_def PosOrd_irrefl by blast |
|
223 |
|
224 lemma PosOrdeq_trans: |
|
225 assumes "v1 :\<sqsubseteq>val v2" "v2 :\<sqsubseteq>val v3" |
|
226 shows "v1 :\<sqsubseteq>val v3" |
|
227 using assms PosOrd_ordering |
|
228 unfolding ordering_def |
|
229 by (metis partial_preordering.trans) |
|
230 |
|
231 lemma PosOrdeq_antisym: |
|
232 assumes "v1 :\<sqsubseteq>val v2" "v2 :\<sqsubseteq>val v1" |
|
233 shows "v1 = v2" |
|
234 using assms PosOrd_ordering |
|
235 unfolding ordering_def |
|
236 by (simp add: ordering_axioms_def) |
|
237 |
|
238 lemma PosOrdeq_refl: |
|
239 shows "v :\<sqsubseteq>val v" |
|
240 unfolding PosOrd_ex_eq_def |
|
241 by auto |
|
242 |
|
243 |
|
244 lemma PosOrd_shorterE: |
|
245 assumes "v1 :\<sqsubset>val v2" |
|
246 shows "length (flat v2) \<le> length (flat v1)" |
|
247 using assms unfolding PosOrd_ex_def PosOrd_def |
|
248 apply(auto) |
|
249 apply(case_tac p) |
|
250 apply(simp add: pflat_len_simps) |
|
251 apply(drule_tac x="[]" in bspec) |
|
252 apply(simp add: Pos_empty) |
|
253 apply(simp add: pflat_len_simps) |
|
254 done |
|
255 |
|
256 lemma PosOrd_shorterI: |
|
257 assumes "length (flat v2) < length (flat v1)" |
|
258 shows "v1 :\<sqsubset>val v2" |
|
259 unfolding PosOrd_ex_def PosOrd_def pflat_len_def |
|
260 using assms Pos_empty by force |
|
261 |
|
262 lemma PosOrd_spreI: |
|
263 assumes "flat v' \<sqsubset>spre flat v" |
|
264 shows "v :\<sqsubset>val v'" |
|
265 using assms |
|
266 apply(rule_tac PosOrd_shorterI) |
|
267 unfolding prefix_list_def sprefix_list_def |
|
268 by (metis append_Nil2 append_eq_conv_conj drop_all le_less_linear) |
|
269 |
|
270 lemma pflat_len_inside: |
|
271 assumes "pflat_len v2 p < pflat_len v1 p" |
|
272 shows "p \<in> Pos v1" |
|
273 using assms |
|
274 unfolding pflat_len_def |
|
275 by (auto split: if_splits) |
|
276 |
|
277 |
|
278 lemma PosOrd_Left_Right: |
|
279 assumes "flat v1 = flat v2" |
|
280 shows "Left v1 :\<sqsubset>val Right v2" |
|
281 unfolding PosOrd_ex_def |
|
282 apply(rule_tac x="[0]" in exI) |
|
283 apply(auto simp add: PosOrd_def pflat_len_simps assms) |
|
284 done |
|
285 |
|
286 lemma PosOrd_LeftE: |
|
287 assumes "Left v1 :\<sqsubset>val Left v2" "flat v1 = flat v2" |
|
288 shows "v1 :\<sqsubset>val v2" |
|
289 using assms |
|
290 unfolding PosOrd_ex_def PosOrd_def2 |
|
291 apply(auto simp add: pflat_len_simps) |
|
292 apply(frule pflat_len_inside) |
|
293 apply(auto simp add: pflat_len_simps) |
|
294 by (metis lex_simps(3) pflat_len_simps(3)) |
|
295 |
|
296 lemma PosOrd_LeftI: |
|
297 assumes "v1 :\<sqsubset>val v2" "flat v1 = flat v2" |
|
298 shows "Left v1 :\<sqsubset>val Left v2" |
|
299 using assms |
|
300 unfolding PosOrd_ex_def PosOrd_def2 |
|
301 apply(auto simp add: pflat_len_simps) |
|
302 by (metis less_numeral_extra(3) lex_simps(3) pflat_len_simps(3)) |
|
303 |
|
304 lemma PosOrd_Left_eq: |
|
305 assumes "flat v1 = flat v2" |
|
306 shows "Left v1 :\<sqsubset>val Left v2 \<longleftrightarrow> v1 :\<sqsubset>val v2" |
|
307 using assms PosOrd_LeftE PosOrd_LeftI |
|
308 by blast |
|
309 |
|
310 |
|
311 lemma PosOrd_RightE: |
|
312 assumes "Right v1 :\<sqsubset>val Right v2" "flat v1 = flat v2" |
|
313 shows "v1 :\<sqsubset>val v2" |
|
314 using assms |
|
315 unfolding PosOrd_ex_def PosOrd_def2 |
|
316 apply(auto simp add: pflat_len_simps) |
|
317 apply(frule pflat_len_inside) |
|
318 apply(auto simp add: pflat_len_simps) |
|
319 by (metis lex_simps(3) pflat_len_simps(5)) |
|
320 |
|
321 lemma PosOrd_RightI: |
|
322 assumes "v1 :\<sqsubset>val v2" "flat v1 = flat v2" |
|
323 shows "Right v1 :\<sqsubset>val Right v2" |
|
324 using assms |
|
325 unfolding PosOrd_ex_def PosOrd_def2 |
|
326 apply(auto simp add: pflat_len_simps) |
|
327 by (metis lex_simps(3) nat_neq_iff pflat_len_simps(5)) |
|
328 |
|
329 |
|
330 lemma PosOrd_Right_eq: |
|
331 assumes "flat v1 = flat v2" |
|
332 shows "Right v1 :\<sqsubset>val Right v2 \<longleftrightarrow> v1 :\<sqsubset>val v2" |
|
333 using assms PosOrd_RightE PosOrd_RightI |
|
334 by blast |
|
335 |
|
336 |
|
337 lemma PosOrd_SeqI1: |
|
338 assumes "v1 :\<sqsubset>val w1" "flat (Seq v1 v2) = flat (Seq w1 w2)" |
|
339 shows "Seq v1 v2 :\<sqsubset>val Seq w1 w2" |
|
340 using assms(1) |
|
341 apply(subst (asm) PosOrd_ex_def) |
|
342 apply(subst (asm) PosOrd_def) |
|
343 apply(clarify) |
|
344 apply(subst PosOrd_ex_def) |
|
345 apply(rule_tac x="0#p" in exI) |
|
346 apply(subst PosOrd_def) |
|
347 apply(rule conjI) |
|
348 apply(simp add: pflat_len_simps) |
|
349 apply(rule ballI) |
|
350 apply(rule impI) |
|
351 apply(simp only: Pos.simps) |
|
352 apply(auto)[1] |
|
353 apply(simp add: pflat_len_simps) |
|
354 apply(auto simp add: pflat_len_simps) |
|
355 using assms(2) |
|
356 apply(simp) |
|
357 apply(metis length_append of_nat_add) |
|
358 done |
|
359 |
|
360 lemma PosOrd_SeqI2: |
|
361 assumes "v2 :\<sqsubset>val w2" "flat v2 = flat w2" |
|
362 shows "Seq v v2 :\<sqsubset>val Seq v w2" |
|
363 using assms(1) |
|
364 apply(subst (asm) PosOrd_ex_def) |
|
365 apply(subst (asm) PosOrd_def) |
|
366 apply(clarify) |
|
367 apply(subst PosOrd_ex_def) |
|
368 apply(rule_tac x="Suc 0#p" in exI) |
|
369 apply(subst PosOrd_def) |
|
370 apply(rule conjI) |
|
371 apply(simp add: pflat_len_simps) |
|
372 apply(rule ballI) |
|
373 apply(rule impI) |
|
374 apply(simp only: Pos.simps) |
|
375 apply(auto)[1] |
|
376 apply(simp add: pflat_len_simps) |
|
377 using assms(2) |
|
378 apply(simp) |
|
379 apply(auto simp add: pflat_len_simps) |
|
380 done |
|
381 |
|
382 lemma PosOrd_Seq_eq: |
|
383 assumes "flat v2 = flat w2" |
|
384 shows "(Seq v v2) :\<sqsubset>val (Seq v w2) \<longleftrightarrow> v2 :\<sqsubset>val w2" |
|
385 using assms |
|
386 apply(auto) |
|
387 prefer 2 |
|
388 apply(simp add: PosOrd_SeqI2) |
|
389 apply(simp add: PosOrd_ex_def) |
|
390 apply(auto) |
|
391 apply(case_tac p) |
|
392 apply(simp add: PosOrd_def pflat_len_simps) |
|
393 apply(case_tac a) |
|
394 apply(simp add: PosOrd_def pflat_len_simps) |
|
395 apply(clarify) |
|
396 apply(case_tac nat) |
|
397 prefer 2 |
|
398 apply(simp add: PosOrd_def pflat_len_simps pflat_len_outside) |
|
399 apply(rule_tac x="list" in exI) |
|
400 apply(auto simp add: PosOrd_def2 pflat_len_simps) |
|
401 apply(smt Collect_disj_eq lex_list.intros(2) mem_Collect_eq pflat_len_simps(2)) |
|
402 apply(smt Collect_disj_eq lex_list.intros(2) mem_Collect_eq pflat_len_simps(2)) |
|
403 done |
|
404 |
|
405 |
|
406 |
|
407 lemma PosOrd_StarsI: |
|
408 assumes "v1 :\<sqsubset>val v2" "flats (v1#vs1) = flats (v2#vs2)" |
|
409 shows "Stars (v1#vs1) :\<sqsubset>val Stars (v2#vs2)" |
|
410 using assms(1) |
|
411 apply(subst (asm) PosOrd_ex_def) |
|
412 apply(subst (asm) PosOrd_def) |
|
413 apply(clarify) |
|
414 apply(subst PosOrd_ex_def) |
|
415 apply(subst PosOrd_def) |
|
416 apply(rule_tac x="0#p" in exI) |
|
417 apply(simp add: pflat_len_Stars_simps pflat_len_simps) |
|
418 using assms(2) |
|
419 apply(simp add: pflat_len_simps) |
|
420 apply(auto simp add: pflat_len_Stars_simps pflat_len_simps) |
|
421 by (metis length_append of_nat_add) |
|
422 |
|
423 lemma PosOrd_StarsI2: |
|
424 assumes "Stars vs1 :\<sqsubset>val Stars vs2" "flats vs1 = flats vs2" |
|
425 shows "Stars (v#vs1) :\<sqsubset>val Stars (v#vs2)" |
|
426 using assms(1) |
|
427 apply(subst (asm) PosOrd_ex_def) |
|
428 apply(subst (asm) PosOrd_def) |
|
429 apply(clarify) |
|
430 apply(subst PosOrd_ex_def) |
|
431 apply(subst PosOrd_def) |
|
432 apply(case_tac p) |
|
433 apply(simp add: pflat_len_simps) |
|
434 apply(rule_tac x="Suc a#list" in exI) |
|
435 apply(auto simp add: pflat_len_Stars_simps pflat_len_simps assms(2)) |
|
436 done |
|
437 |
|
438 lemma PosOrd_Stars_appendI: |
|
439 assumes "Stars vs1 :\<sqsubset>val Stars vs2" "flat (Stars vs1) = flat (Stars vs2)" |
|
440 shows "Stars (vs @ vs1) :\<sqsubset>val Stars (vs @ vs2)" |
|
441 using assms |
|
442 apply(induct vs) |
|
443 apply(simp) |
|
444 apply(simp add: PosOrd_StarsI2) |
|
445 done |
|
446 |
|
447 lemma PosOrd_StarsE2: |
|
448 assumes "Stars (v # vs1) :\<sqsubset>val Stars (v # vs2)" |
|
449 shows "Stars vs1 :\<sqsubset>val Stars vs2" |
|
450 using assms |
|
451 apply(subst (asm) PosOrd_ex_def) |
|
452 apply(erule exE) |
|
453 apply(case_tac p) |
|
454 apply(simp) |
|
455 apply(simp add: PosOrd_def pflat_len_simps) |
|
456 apply(subst PosOrd_ex_def) |
|
457 apply(rule_tac x="[]" in exI) |
|
458 apply(simp add: PosOrd_def pflat_len_simps Pos_empty) |
|
459 apply(simp) |
|
460 apply(case_tac a) |
|
461 apply(clarify) |
|
462 apply(auto simp add: pflat_len_simps PosOrd_def pflat_len_def split: if_splits)[1] |
|
463 apply(clarify) |
|
464 apply(simp add: PosOrd_ex_def) |
|
465 apply(rule_tac x="nat#list" in exI) |
|
466 apply(auto simp add: PosOrd_def pflat_len_simps)[1] |
|
467 apply(case_tac q) |
|
468 apply(simp add: PosOrd_def pflat_len_simps) |
|
469 apply(clarify) |
|
470 apply(drule_tac x="Suc a # lista" in bspec) |
|
471 apply(simp) |
|
472 apply(auto simp add: PosOrd_def pflat_len_simps)[1] |
|
473 apply(case_tac q) |
|
474 apply(simp add: PosOrd_def pflat_len_simps) |
|
475 apply(clarify) |
|
476 apply(drule_tac x="Suc a # lista" in bspec) |
|
477 apply(simp) |
|
478 apply(auto simp add: PosOrd_def pflat_len_simps)[1] |
|
479 done |
|
480 |
|
481 lemma PosOrd_Stars_appendE: |
|
482 assumes "Stars (vs @ vs1) :\<sqsubset>val Stars (vs @ vs2)" |
|
483 shows "Stars vs1 :\<sqsubset>val Stars vs2" |
|
484 using assms |
|
485 apply(induct vs) |
|
486 apply(simp) |
|
487 apply(simp add: PosOrd_StarsE2) |
|
488 done |
|
489 |
|
490 lemma PosOrd_Stars_append_eq: |
|
491 assumes "flats vs1 = flats vs2" |
|
492 shows "Stars (vs @ vs1) :\<sqsubset>val Stars (vs @ vs2) \<longleftrightarrow> Stars vs1 :\<sqsubset>val Stars vs2" |
|
493 using assms |
|
494 apply(rule_tac iffI) |
|
495 apply(erule PosOrd_Stars_appendE) |
|
496 apply(rule PosOrd_Stars_appendI) |
|
497 apply(auto) |
|
498 done |
|
499 |
|
500 lemma PosOrd_almost_trichotomous: |
|
501 shows "v1 :\<sqsubset>val v2 \<or> v2 :\<sqsubset>val v1 \<or> (length (flat v1) = length (flat v2))" |
|
502 apply(auto simp add: PosOrd_ex_def) |
|
503 apply(auto simp add: PosOrd_def) |
|
504 apply(rule_tac x="[]" in exI) |
|
505 apply(auto simp add: Pos_empty pflat_len_simps) |
|
506 apply(drule_tac x="[]" in spec) |
|
507 apply(auto simp add: Pos_empty pflat_len_simps) |
|
508 done |
|
509 |
|
510 |
|
511 |
|
512 section \<open>The Posix Value is smaller than any other Value\<close> |
|
513 |
|
514 |
|
515 lemma Posix_PosOrd: |
|
516 assumes "s \<in> r \<rightarrow> v1" "v2 \<in> LV r s" |
|
517 shows "v1 :\<sqsubseteq>val v2" |
|
518 using assms |
|
519 proof (induct arbitrary: v2 rule: Posix.induct) |
|
520 case (Posix_ONE v) |
|
521 have "v \<in> LV ONE []" by fact |
|
522 then have "v = Void" |
|
523 by (simp add: LV_simps) |
|
524 then show "Void :\<sqsubseteq>val v" |
|
525 by (simp add: PosOrd_ex_eq_def) |
|
526 next |
|
527 case (Posix_CH c v) |
|
528 have "v \<in> LV (CH c) [c]" by fact |
|
529 then have "v = Char c" |
|
530 by (simp add: LV_simps) |
|
531 then show "Char c :\<sqsubseteq>val v" |
|
532 by (simp add: PosOrd_ex_eq_def) |
|
533 next |
|
534 case (Posix_ALT1 s r1 v r2 v2) |
|
535 have as1: "s \<in> r1 \<rightarrow> v" by fact |
|
536 have IH: "\<And>v2. v2 \<in> LV r1 s \<Longrightarrow> v :\<sqsubseteq>val v2" by fact |
|
537 have "v2 \<in> LV (ALT r1 r2) s" by fact |
|
538 then have "\<Turnstile> v2 : ALT r1 r2" "flat v2 = s" |
|
539 by(auto simp add: LV_def prefix_list_def) |
|
540 then consider |
|
541 (Left) v3 where "v2 = Left v3" "\<Turnstile> v3 : r1" "flat v3 = s" |
|
542 | (Right) v3 where "v2 = Right v3" "\<Turnstile> v3 : r2" "flat v3 = s" |
|
543 by (auto elim: Prf.cases) |
|
544 then show "Left v :\<sqsubseteq>val v2" |
|
545 proof(cases) |
|
546 case (Left v3) |
|
547 have "v3 \<in> LV r1 s" using Left(2,3) |
|
548 by (auto simp add: LV_def prefix_list_def) |
|
549 with IH have "v :\<sqsubseteq>val v3" by simp |
|
550 moreover |
|
551 have "flat v3 = flat v" using as1 Left(3) |
|
552 by (simp add: Posix1(2)) |
|
553 ultimately have "Left v :\<sqsubseteq>val Left v3" |
|
554 by (simp add: PosOrd_ex_eq_def PosOrd_Left_eq) |
|
555 then show "Left v :\<sqsubseteq>val v2" unfolding Left . |
|
556 next |
|
557 case (Right v3) |
|
558 have "flat v3 = flat v" using as1 Right(3) |
|
559 by (simp add: Posix1(2)) |
|
560 then have "Left v :\<sqsubseteq>val Right v3" |
|
561 unfolding PosOrd_ex_eq_def |
|
562 by (simp add: PosOrd_Left_Right) |
|
563 then show "Left v :\<sqsubseteq>val v2" unfolding Right . |
|
564 qed |
|
565 next |
|
566 case (Posix_ALT2 s r2 v r1 v2) |
|
567 have as1: "s \<in> r2 \<rightarrow> v" by fact |
|
568 have as2: "s \<notin> L r1" by fact |
|
569 have IH: "\<And>v2. v2 \<in> LV r2 s \<Longrightarrow> v :\<sqsubseteq>val v2" by fact |
|
570 have "v2 \<in> LV (ALT r1 r2) s" by fact |
|
571 then have "\<Turnstile> v2 : ALT r1 r2" "flat v2 = s" |
|
572 by(auto simp add: LV_def prefix_list_def) |
|
573 then consider |
|
574 (Left) v3 where "v2 = Left v3" "\<Turnstile> v3 : r1" "flat v3 = s" |
|
575 | (Right) v3 where "v2 = Right v3" "\<Turnstile> v3 : r2" "flat v3 = s" |
|
576 by (auto elim: Prf.cases) |
|
577 then show "Right v :\<sqsubseteq>val v2" |
|
578 proof (cases) |
|
579 case (Right v3) |
|
580 have "v3 \<in> LV r2 s" using Right(2,3) |
|
581 by (auto simp add: LV_def prefix_list_def) |
|
582 with IH have "v :\<sqsubseteq>val v3" by simp |
|
583 moreover |
|
584 have "flat v3 = flat v" using as1 Right(3) |
|
585 by (simp add: Posix1(2)) |
|
586 ultimately have "Right v :\<sqsubseteq>val Right v3" |
|
587 by (auto simp add: PosOrd_ex_eq_def PosOrd_RightI) |
|
588 then show "Right v :\<sqsubseteq>val v2" unfolding Right . |
|
589 next |
|
590 case (Left v3) |
|
591 have "v3 \<in> LV r1 s" using Left(2,3) as2 |
|
592 by (auto simp add: LV_def prefix_list_def) |
|
593 then have "flat v3 = flat v \<and> \<Turnstile> v3 : r1" using as1 Left(3) |
|
594 by (simp add: Posix1(2) LV_def) |
|
595 then have "False" using as1 as2 Left |
|
596 by (auto simp add: Posix1(2) L_flat_Prf1) |
|
597 then show "Right v :\<sqsubseteq>val v2" by simp |
|
598 qed |
|
599 next |
|
600 case (Posix_SEQ s1 r1 v1 s2 r2 v2 v3) |
|
601 have "s1 \<in> r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" by fact+ |
|
602 then have as1: "s1 = flat v1" "s2 = flat v2" by (simp_all add: Posix1(2)) |
|
603 have IH1: "\<And>v3. v3 \<in> LV r1 s1 \<Longrightarrow> v1 :\<sqsubseteq>val v3" by fact |
|
604 have IH2: "\<And>v3. v3 \<in> LV r2 s2 \<Longrightarrow> v2 :\<sqsubseteq>val v3" by fact |
|
605 have cond: "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by fact |
|
606 have "v3 \<in> LV (SEQ r1 r2) (s1 @ s2)" by fact |
|
607 then obtain v3a v3b where eqs: |
|
608 "v3 = Seq v3a v3b" "\<Turnstile> v3a : r1" "\<Turnstile> v3b : r2" |
|
609 "flat v3a @ flat v3b = s1 @ s2" |
|
610 by (force simp add: prefix_list_def LV_def elim: Prf.cases) |
|
611 with cond have "flat v3a \<sqsubseteq>pre s1" unfolding prefix_list_def |
|
612 by (smt L_flat_Prf1 append_eq_append_conv2 append_self_conv) |
|
613 then have "flat v3a \<sqsubset>spre s1 \<or> (flat v3a = s1 \<and> flat v3b = s2)" using eqs |
|
614 by (simp add: sprefix_list_def append_eq_conv_conj) |
|
615 then have q2: "v1 :\<sqsubset>val v3a \<or> (flat v3a = s1 \<and> flat v3b = s2)" |
|
616 using PosOrd_spreI as1(1) eqs by blast |
|
617 then have "v1 :\<sqsubset>val v3a \<or> (v3a \<in> LV r1 s1 \<and> v3b \<in> LV r2 s2)" using eqs(2,3) |
|
618 by (auto simp add: LV_def) |
|
619 then have "v1 :\<sqsubset>val v3a \<or> (v1 :\<sqsubseteq>val v3a \<and> v2 :\<sqsubseteq>val v3b)" using IH1 IH2 by blast |
|
620 then have "Seq v1 v2 :\<sqsubseteq>val Seq v3a v3b" using eqs q2 as1 |
|
621 unfolding PosOrd_ex_eq_def by (auto simp add: PosOrd_SeqI1 PosOrd_Seq_eq) |
|
622 then show "Seq v1 v2 :\<sqsubseteq>val v3" unfolding eqs by blast |
|
623 next |
|
624 case (Posix_STAR1 s1 r v s2 vs v3) |
|
625 have "s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" by fact+ |
|
626 then have as1: "s1 = flat v" "s2 = flat (Stars vs)" by (auto dest: Posix1(2)) |
|
627 have IH1: "\<And>v3. v3 \<in> LV r s1 \<Longrightarrow> v :\<sqsubseteq>val v3" by fact |
|
628 have IH2: "\<And>v3. v3 \<in> LV (STAR r) s2 \<Longrightarrow> Stars vs :\<sqsubseteq>val v3" by fact |
|
629 have cond: "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" by fact |
|
630 have cond2: "flat v \<noteq> []" by fact |
|
631 have "v3 \<in> LV (STAR r) (s1 @ s2)" by fact |
|
632 then consider |
|
633 (NonEmpty) v3a vs3 where "v3 = Stars (v3a # vs3)" |
|
634 "\<Turnstile> v3a : r" "\<Turnstile> Stars vs3 : STAR r" |
|
635 "flat (Stars (v3a # vs3)) = s1 @ s2" |
|
636 | (Empty) "v3 = Stars []" |
|
637 unfolding LV_def |
|
638 apply(auto) |
|
639 apply(erule Prf.cases) |
|
640 apply(auto) |
|
641 apply(case_tac vs) |
|
642 apply(auto intro: Prf.intros) |
|
643 done |
|
644 then show "Stars (v # vs) :\<sqsubseteq>val v3" |
|
645 proof (cases) |
|
646 case (NonEmpty v3a vs3) |
|
647 have "flat (Stars (v3a # vs3)) = s1 @ s2" using NonEmpty(4) . |
|
648 with cond have "flat v3a \<sqsubseteq>pre s1" using NonEmpty(2,3) |
|
649 unfolding prefix_list_def |
|
650 by (smt L_flat_Prf1 append_Nil2 append_eq_append_conv2 flat.simps(7)) |
|
651 then have "flat v3a \<sqsubset>spre s1 \<or> (flat v3a = s1 \<and> flat (Stars vs3) = s2)" using NonEmpty(4) |
|
652 by (simp add: sprefix_list_def append_eq_conv_conj) |
|
653 then have q2: "v :\<sqsubset>val v3a \<or> (flat v3a = s1 \<and> flat (Stars vs3) = s2)" |
|
654 using PosOrd_spreI as1(1) NonEmpty(4) by blast |
|
655 then have "v :\<sqsubset>val v3a \<or> (v3a \<in> LV r s1 \<and> Stars vs3 \<in> LV (STAR r) s2)" |
|
656 using NonEmpty(2,3) by (auto simp add: LV_def) |
|
657 then have "v :\<sqsubset>val v3a \<or> (v :\<sqsubseteq>val v3a \<and> Stars vs :\<sqsubseteq>val Stars vs3)" using IH1 IH2 by blast |
|
658 then have "v :\<sqsubset>val v3a \<or> (v = v3a \<and> Stars vs :\<sqsubseteq>val Stars vs3)" |
|
659 unfolding PosOrd_ex_eq_def by auto |
|
660 then have "Stars (v # vs) :\<sqsubseteq>val Stars (v3a # vs3)" using NonEmpty(4) q2 as1 |
|
661 unfolding PosOrd_ex_eq_def |
|
662 using PosOrd_StarsI PosOrd_StarsI2 by auto |
|
663 then show "Stars (v # vs) :\<sqsubseteq>val v3" unfolding NonEmpty by blast |
|
664 next |
|
665 case Empty |
|
666 have "v3 = Stars []" by fact |
|
667 then show "Stars (v # vs) :\<sqsubseteq>val v3" |
|
668 unfolding PosOrd_ex_eq_def using cond2 |
|
669 by (simp add: PosOrd_shorterI) |
|
670 qed |
|
671 next |
|
672 case (Posix_STAR2 r v2) |
|
673 have "v2 \<in> LV (STAR r) []" by fact |
|
674 then have "v2 = Stars []" |
|
675 unfolding LV_def by (auto elim: Prf.cases) |
|
676 then show "Stars [] :\<sqsubseteq>val v2" |
|
677 by (simp add: PosOrd_ex_eq_def) |
|
678 qed |
|
679 |
|
680 |
|
681 lemma Posix_PosOrd_reverse: |
|
682 assumes "s \<in> r \<rightarrow> v1" |
|
683 shows "\<not>(\<exists>v2 \<in> LV r s. v2 :\<sqsubset>val v1)" |
|
684 using assms |
|
685 by (metis Posix_PosOrd less_irrefl PosOrd_def |
|
686 PosOrd_ex_eq_def PosOrd_ex_def PosOrd_trans) |
|
687 |
|
688 lemma PosOrd_Posix: |
|
689 assumes "v1 \<in> LV r s" "\<forall>v\<^sub>2 \<in> LV r s. \<not> v\<^sub>2 :\<sqsubset>val v1" |
|
690 shows "s \<in> r \<rightarrow> v1" |
|
691 proof - |
|
692 have "s \<in> L r" using assms(1) unfolding LV_def |
|
693 using L_flat_Prf1 by blast |
|
694 then obtain vposix where vp: "s \<in> r \<rightarrow> vposix" |
|
695 using lexer_correct_Some by blast |
|
696 with assms(1) have "vposix :\<sqsubseteq>val v1" by (simp add: Posix_PosOrd) |
|
697 then have "vposix = v1 \<or> vposix :\<sqsubset>val v1" unfolding PosOrd_ex_eq2 by auto |
|
698 moreover |
|
699 { assume "vposix :\<sqsubset>val v1" |
|
700 moreover |
|
701 have "vposix \<in> LV r s" using vp |
|
702 using Posix_LV by blast |
|
703 ultimately have "False" using assms(2) by blast |
|
704 } |
|
705 ultimately show "s \<in> r \<rightarrow> v1" using vp by blast |
|
706 qed |
|
707 |
|
708 lemma Least_existence: |
|
709 assumes "LV r s \<noteq> {}" |
|
710 shows " \<exists>vmin \<in> LV r s. \<forall>v \<in> LV r s. vmin :\<sqsubseteq>val v" |
|
711 proof - |
|
712 from assms |
|
713 obtain vposix where "s \<in> r \<rightarrow> vposix" |
|
714 unfolding LV_def |
|
715 using L_flat_Prf1 lexer_correct_Some by blast |
|
716 then have "\<forall>v \<in> LV r s. vposix :\<sqsubseteq>val v" |
|
717 by (simp add: Posix_PosOrd) |
|
718 then show "\<exists>vmin \<in> LV r s. \<forall>v \<in> LV r s. vmin :\<sqsubseteq>val v" |
|
719 using Posix_LV \<open>s \<in> r \<rightarrow> vposix\<close> by blast |
|
720 qed |
|
721 |
|
722 lemma Least_existence1: |
|
723 assumes "LV r s \<noteq> {}" |
|
724 shows " \<exists>!vmin \<in> LV r s. \<forall>v \<in> LV r s. vmin :\<sqsubseteq>val v" |
|
725 using Least_existence[OF assms] assms |
|
726 using PosOrdeq_antisym by blast |
|
727 |
|
728 lemma Least_existence2: |
|
729 assumes "LV r s \<noteq> {}" |
|
730 shows " \<exists>!vmin \<in> LV r s. lexer r s = Some vmin \<and> (\<forall>v \<in> LV r s. vmin :\<sqsubseteq>val v)" |
|
731 using Least_existence[OF assms] assms |
|
732 using PosOrdeq_antisym |
|
733 using PosOrd_Posix PosOrd_ex_eq2 lexer_correctness(1) by auto |
|
734 |
|
735 |
|
736 lemma Least_existence1_pre: |
|
737 assumes "LV r s \<noteq> {}" |
|
738 shows " \<exists>!vmin \<in> LV r s. \<forall>v \<in> (LV r s \<union> {v'. flat v' \<sqsubset>spre s}). vmin :\<sqsubseteq>val v" |
|
739 using Least_existence[OF assms] assms |
|
740 apply - |
|
741 apply(erule bexE) |
|
742 apply(rule_tac a="vmin" in ex1I) |
|
743 apply(auto)[1] |
|
744 apply (metis PosOrd_Posix PosOrd_ex_eq2 PosOrd_spreI PosOrdeq_antisym Posix1(2)) |
|
745 apply(auto)[1] |
|
746 apply(simp add: PosOrdeq_antisym) |
|
747 done |
|
748 |
|
749 lemma |
|
750 shows "partial_order_on UNIV {(v1, v2). v1 :\<sqsubseteq>val v2}" |
|
751 apply(simp add: partial_order_on_def) |
|
752 apply(simp add: preorder_on_def refl_on_def) |
|
753 apply(simp add: PosOrdeq_refl) |
|
754 apply(auto) |
|
755 apply(rule transI) |
|
756 apply(auto intro: PosOrdeq_trans)[1] |
|
757 apply(rule antisymI) |
|
758 apply(simp add: PosOrdeq_antisym) |
|
759 done |
|
760 |
|
761 lemma |
|
762 "wf {(v1, v2). v1 :\<sqsubset>val v2 \<and> v1 \<in> LV r s \<and> v2 \<in> LV r s}" |
|
763 apply(rule finite_acyclic_wf) |
|
764 prefer 2 |
|
765 apply(simp add: acyclic_def) |
|
766 apply(induct_tac rule: trancl.induct) |
|
767 apply(auto)[1] |
|
768 oops |
|
769 |
|
770 |
|
771 unused_thms |
|
772 |
|
773 end |