author | Christian Urban <urbanc@in.tum.de> |
Tue, 04 Jul 2017 18:09:29 +0100 | |
changeset 263 | 00c9a95d492e |
parent 262 | 45ad887fa6aa |
child 264 | e2828c4a1e23 |
permissions | -rw-r--r-- |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
theory Positions |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
imports "Lexer" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
begin |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
|
254 | 6 |
|
7 |
section {* Position in values *} |
|
8 |
||
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
fun |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
at :: "val \<Rightarrow> nat list \<Rightarrow> val" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
where |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
"at v [] = v" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
| "at (Left v) (0#ps)= at v ps" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
| "at (Right v) (Suc 0#ps)= at v ps" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
| "at (Seq v1 v2) (0#ps)= at v1 ps" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
| "at (Seq v1 v2) (Suc 0#ps)= at v2 ps" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
| "at (Stars vs) (n#ps)= at (nth vs n) ps" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
fun Pos :: "val \<Rightarrow> (nat list) set" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
where |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
"Pos (Void) = {[]}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
| "Pos (Char c) = {[]}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
| "Pos (Left v) = {[]} \<union> {0#ps | ps. ps \<in> Pos v}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
| "Pos (Right v) = {[]} \<union> {1#ps | ps. ps \<in> Pos v}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
| "Pos (Seq v1 v2) = {[]} \<union> {0#ps | ps. ps \<in> Pos v1} \<union> {1#ps | ps. ps \<in> Pos v2}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
| "Pos (Stars []) = {[]}" |
251 | 27 |
| "Pos (Stars (v#vs)) = {[]} \<union> {0#ps | ps. ps \<in> Pos v} \<union> {Suc n#ps | n ps. n#ps \<in> Pos (Stars vs)}" |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
|
253 | 29 |
lemma Pos_stars: |
30 |
"Pos (Stars vs) = {[]} \<union> (\<Union>n < length vs. {n#ps | ps. ps \<in> Pos (vs ! n)})" |
|
31 |
apply(induct vs) |
|
32 |
apply(simp) |
|
33 |
apply(simp) |
|
34 |
apply(simp add: insert_ident) |
|
35 |
apply(rule subset_antisym) |
|
36 |
apply(auto) |
|
37 |
apply(auto)[1] |
|
38 |
using less_Suc_eq_0_disj by auto |
|
39 |
||
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
lemma Pos_empty: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
41 |
shows "[] \<in> Pos v" |
251 | 42 |
by (induct v rule: Pos.induct)(auto) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
fun intlen :: "'a list \<Rightarrow> int" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
45 |
where |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
46 |
"intlen [] = 0" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
47 |
| "intlen (x#xs) = 1 + intlen xs" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
|
255 | 49 |
lemma intlen_bigger: |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
50 |
shows "0 \<le> intlen xs" |
251 | 51 |
by (induct xs)(auto) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
52 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
lemma intlen_append: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
shows "intlen (xs @ ys) = intlen xs + intlen ys" |
251 | 55 |
by (induct xs arbitrary: ys) (auto) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
lemma intlen_length: |
256 | 58 |
shows "intlen xs < intlen ys \<longleftrightarrow> length xs < length ys" |
59 |
apply(induct xs arbitrary: ys) |
|
60 |
apply(auto) |
|
61 |
apply(case_tac ys) |
|
62 |
apply(simp_all) |
|
63 |
apply (smt intlen_bigger) |
|
257 | 64 |
apply (smt Suc_mono intlen.simps(2) length_Suc_conv less_antisym) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
65 |
by (smt Suc_lessE intlen.simps(2) length_Suc_conv) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
66 |
|
255 | 67 |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
68 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
69 |
definition pflat_len :: "val \<Rightarrow> nat list => int" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
70 |
where |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
71 |
"pflat_len v p \<equiv> (if p \<in> Pos v then intlen (flat (at v p)) else -1)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
72 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
lemma pflat_len_simps: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
shows "pflat_len (Seq v1 v2) (0#p) = pflat_len v1 p" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
and "pflat_len (Seq v1 v2) (Suc 0#p) = pflat_len v2 p" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
and "pflat_len (Left v) (0#p) = pflat_len v p" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
77 |
and "pflat_len (Left v) (Suc 0#p) = -1" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
78 |
and "pflat_len (Right v) (Suc 0#p) = pflat_len v p" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
79 |
and "pflat_len (Right v) (0#p) = -1" |
251 | 80 |
and "pflat_len (Stars (v#vs)) (Suc n#p) = pflat_len (Stars vs) (n#p)" |
81 |
and "pflat_len (Stars (v#vs)) (0#p) = pflat_len v p" |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
82 |
and "pflat_len v [] = intlen (flat v)" |
251 | 83 |
by (auto simp add: pflat_len_def Pos_empty) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
84 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
85 |
lemma pflat_len_Stars_simps: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
86 |
assumes "n < length vs" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
87 |
shows "pflat_len (Stars vs) (n#p) = pflat_len (vs!n) p" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
88 |
using assms |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
89 |
apply(induct vs arbitrary: n p) |
251 | 90 |
apply(auto simp add: less_Suc_eq_0_disj pflat_len_simps) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
91 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
92 |
|
254 | 93 |
lemma outside_lemma: |
94 |
assumes "p \<notin> Pos v1 \<union> Pos v2" |
|
95 |
shows "pflat_len v1 p = pflat_len v2 p" |
|
96 |
using assms by (auto simp add: pflat_len_def) |
|
252 | 97 |
|
98 |
||
99 |
section {* Orderings *} |
|
100 |
||
101 |
||
257 | 102 |
definition prefix_list:: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<sqsubseteq>pre _" [60,59] 60) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
103 |
where |
252 | 104 |
"ps1 \<sqsubseteq>pre ps2 \<equiv> \<exists>ps'. ps1 @ps' = ps2" |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
105 |
|
257 | 106 |
definition sprefix_list:: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<sqsubset>spre _" [60,59] 60) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
107 |
where |
252 | 108 |
"ps1 \<sqsubset>spre ps2 \<equiv> ps1 \<sqsubseteq>pre ps2 \<and> ps1 \<noteq> ps2" |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
109 |
|
257 | 110 |
inductive lex_list :: "nat list \<Rightarrow> nat list \<Rightarrow> bool" ("_ \<sqsubset>lex _" [60,59] 60) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
111 |
where |
252 | 112 |
"[] \<sqsubset>lex (p#ps)" |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
113 |
| "ps1 \<sqsubset>lex ps2 \<Longrightarrow> (p#ps1) \<sqsubset>lex (p#ps2)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
114 |
| "p1 < p2 \<Longrightarrow> (p1#ps1) \<sqsubset>lex (p2#ps2)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
115 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
116 |
lemma lex_irrfl: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
117 |
fixes ps1 ps2 :: "nat list" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
118 |
assumes "ps1 \<sqsubset>lex ps2" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
119 |
shows "ps1 \<noteq> ps2" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
120 |
using assms |
257 | 121 |
by(induct rule: lex_list.induct)(auto) |
122 |
||
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
123 |
|
251 | 124 |
lemma lex_simps [simp]: |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
125 |
fixes xs ys :: "nat list" |
251 | 126 |
shows "[] \<sqsubset>lex ys \<longleftrightarrow> ys \<noteq> []" |
127 |
and "xs \<sqsubset>lex [] \<longleftrightarrow> False" |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
128 |
and "(x # xs) \<sqsubset>lex (y # ys) \<longleftrightarrow> (x < y \<or> (\<not> y < x \<and> xs \<sqsubset>lex ys))" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
129 |
apply - |
251 | 130 |
apply (metis lex_irrfl lex_list.intros(1) list.exhaust) |
131 |
using lex_list.cases apply blast |
|
132 |
using lex_list.cases lex_list.intros(2) lex_list.intros(3) not_less_iff_gr_or_eq by fastforce |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
133 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
134 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
135 |
lemma lex_trans: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
136 |
fixes ps1 ps2 ps3 :: "nat list" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
137 |
assumes "ps1 \<sqsubset>lex ps2" "ps2 \<sqsubset>lex ps3" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
138 |
shows "ps1 \<sqsubset>lex ps3" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
139 |
using assms |
251 | 140 |
apply(induct arbitrary: ps3 rule: lex_list.induct) |
141 |
apply(erule lex_list.cases) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
142 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
143 |
apply(rotate_tac 2) |
251 | 144 |
apply(erule lex_list.cases) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
145 |
apply(simp_all) |
251 | 146 |
apply(erule lex_list.cases) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
147 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
148 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
149 |
|
257 | 150 |
lemma lex_trichotomous: |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
151 |
fixes p q :: "nat list" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
152 |
shows "p = q \<or> p \<sqsubset>lex q \<or> q \<sqsubset>lex p" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
153 |
apply(induct p arbitrary: q) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
154 |
apply(auto) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
155 |
apply(case_tac q) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
156 |
apply(auto) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
157 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
158 |
|
254 | 159 |
|
160 |
||
161 |
||
162 |
section {* Ordering of values according to Okui & Suzuki *} |
|
163 |
||
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
164 |
|
261 | 165 |
definition PosOrd:: "val \<Rightarrow> nat list \<Rightarrow> val \<Rightarrow> bool" ("_ \<sqsubset>val _ _" [60, 60, 59] 60) |
166 |
where |
|
167 |
"v1 \<sqsubset>val p v2 \<equiv> p \<in> Pos v1 \<and> |
|
168 |
pflat_len v1 p > pflat_len v2 p \<and> |
|
169 |
(\<forall>q \<in> Pos v1 \<union> Pos v2. q \<sqsubset>lex p \<longrightarrow> pflat_len v1 q = pflat_len v2 q)" |
|
170 |
||
171 |
definition ValFlat_ord:: "val \<Rightarrow> nat list \<Rightarrow> val \<Rightarrow> bool" ("_ \<sqsubset>fval _ _") |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
172 |
where |
261 | 173 |
"v1 \<sqsubset>fval p v2 \<equiv> p \<in> Pos v1 \<and> |
174 |
(p \<notin> Pos v2 \<or> flat (at v2 p) \<sqsubset>spre flat (at v1 p)) \<and> |
|
175 |
(\<forall>q \<in> Pos v1 \<union> Pos v2. q \<sqsubset>lex p \<longrightarrow> (pflat_len v1 q = pflat_len v2 q))" |
|
176 |
||
177 |
lemma |
|
178 |
assumes "v1 \<sqsubset>fval p v2" |
|
179 |
shows "v1 \<sqsubset>val p v2" |
|
180 |
using assms |
|
181 |
unfolding ValFlat_ord_def PosOrd_def |
|
182 |
apply(clarify) |
|
183 |
apply(simp add: pflat_len_def) |
|
184 |
apply(auto)[1] |
|
185 |
apply(smt intlen_bigger) |
|
186 |
apply(simp add: sprefix_list_def prefix_list_def) |
|
187 |
apply(auto)[1] |
|
188 |
apply(drule sym) |
|
189 |
apply(simp add: intlen_append) |
|
190 |
apply (metis intlen.simps(1) intlen_length length_greater_0_conv list.size(3)) |
|
191 |
apply(smt intlen_bigger) |
|
192 |
done |
|
193 |
||
194 |
lemma |
|
195 |
assumes "v1 \<sqsubset>val p v2" "flat (at v2 p) \<sqsubset>spre flat (at v1 p)" |
|
196 |
shows "v1 \<sqsubset>fval p v2" |
|
197 |
using assms |
|
198 |
unfolding ValFlat_ord_def PosOrd_def |
|
199 |
apply(clarify) |
|
200 |
done |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
201 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
202 |
|
261 | 203 |
definition PosOrd_ex:: "val \<Rightarrow> val \<Rightarrow> bool" ("_ :\<sqsubset>val _" [60, 59] 60) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
204 |
where |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
205 |
"v1 :\<sqsubset>val v2 \<equiv> (\<exists>p. v1 \<sqsubset>val p v2)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
206 |
|
261 | 207 |
definition PosOrd_ex1:: "val \<Rightarrow> val \<Rightarrow> bool" ("_ :\<sqsubseteq>val _" [60, 59] 60) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
208 |
where |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
209 |
"v1 :\<sqsubseteq>val v2 \<equiv> v1 :\<sqsubset>val v2 \<or> v1 = v2" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
210 |
|
261 | 211 |
lemma PosOrd_shorterE: |
255 | 212 |
assumes "v1 :\<sqsubset>val v2" |
213 |
shows "length (flat v2) \<le> length (flat v1)" |
|
214 |
using assms |
|
261 | 215 |
apply(auto simp add: PosOrd_ex_def PosOrd_def) |
255 | 216 |
apply(case_tac p) |
217 |
apply(simp add: pflat_len_simps) |
|
257 | 218 |
apply(simp add: intlen_length) |
255 | 219 |
apply(simp) |
220 |
apply(drule_tac x="[]" in bspec) |
|
221 |
apply(simp add: Pos_empty) |
|
222 |
apply(simp add: pflat_len_simps) |
|
223 |
by (metis intlen_length le_less less_irrefl linear) |
|
224 |
||
225 |
||
261 | 226 |
lemma PosOrd_shorterI: |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
227 |
assumes "length (flat v') < length (flat v)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
228 |
shows "v :\<sqsubset>val v'" |
251 | 229 |
using assms |
261 | 230 |
unfolding PosOrd_ex_def |
231 |
by (metis Pos_empty intlen_length lex_simps(2) pflat_len_simps(9) PosOrd_def) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
232 |
|
261 | 233 |
lemma PosOrd_spreI: |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
234 |
assumes "(flat v') \<sqsubset>spre (flat v)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
235 |
shows "v :\<sqsubset>val v'" |
251 | 236 |
using assms |
261 | 237 |
apply(rule_tac PosOrd_shorterI) |
251 | 238 |
by (metis append_eq_conv_conj le_less_linear prefix_list_def sprefix_list_def take_all) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
239 |
|
261 | 240 |
lemma PosOrd_Left_Right: |
241 |
assumes "flat v1 = flat v2" |
|
242 |
shows "Left v1 :\<sqsubset>val Right v2" |
|
243 |
unfolding PosOrd_ex_def |
|
244 |
apply(rule_tac x="[0]" in exI) |
|
245 |
using assms |
|
246 |
apply(auto simp add: PosOrd_def pflat_len_simps Pos_empty) |
|
247 |
apply(smt intlen_bigger) |
|
248 |
done |
|
256 | 249 |
|
261 | 250 |
lemma PosOrd_LeftI: |
251 | 251 |
assumes "v :\<sqsubset>val v'" "flat v = flat v'" |
252 |
shows "(Left v) :\<sqsubset>val (Left v')" |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
253 |
using assms(1) |
261 | 254 |
unfolding PosOrd_ex_def |
251 | 255 |
apply(auto) |
256 |
apply(rule_tac x="0#p" in exI) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
257 |
using assms(2) |
261 | 258 |
apply(auto simp add: PosOrd_def pflat_len_simps) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
259 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
260 |
|
261 | 261 |
lemma PosOrd_RightI: |
251 | 262 |
assumes "v :\<sqsubset>val v'" "flat v = flat v'" |
263 |
shows "(Right v) :\<sqsubset>val (Right v')" |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
264 |
using assms(1) |
261 | 265 |
unfolding PosOrd_ex_def |
251 | 266 |
apply(auto) |
267 |
apply(rule_tac x="Suc 0#p" in exI) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
268 |
using assms(2) |
261 | 269 |
apply(auto simp add: PosOrd_def pflat_len_simps) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
270 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
271 |
|
261 | 272 |
lemma PosOrd_LeftE: |
252 | 273 |
assumes "(Left v1) :\<sqsubset>val (Left v2)" |
274 |
shows "v1 :\<sqsubset>val v2" |
|
275 |
using assms |
|
261 | 276 |
apply(simp add: PosOrd_ex_def) |
252 | 277 |
apply(erule exE) |
278 |
apply(case_tac "p = []") |
|
261 | 279 |
apply(simp add: PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
280 |
apply(auto simp add: pflat_len_simps) |
252 | 281 |
apply(rule_tac x="[]" in exI) |
282 |
apply(simp add: Pos_empty pflat_len_simps) |
|
261 | 283 |
apply(auto simp add: pflat_len_simps PosOrd_def) |
252 | 284 |
apply(rule_tac x="ps" in exI) |
285 |
apply(auto) |
|
286 |
apply(drule_tac x="0#q" in bspec) |
|
287 |
apply(simp) |
|
288 |
apply(simp add: pflat_len_simps) |
|
289 |
apply(drule_tac x="0#q" in bspec) |
|
290 |
apply(simp) |
|
291 |
apply(simp add: pflat_len_simps) |
|
292 |
done |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
293 |
|
261 | 294 |
lemma PosOrd_RightE: |
252 | 295 |
assumes "(Right v1) :\<sqsubset>val (Right v2)" |
296 |
shows "v1 :\<sqsubset>val v2" |
|
297 |
using assms |
|
261 | 298 |
apply(simp add: PosOrd_ex_def) |
252 | 299 |
apply(erule exE) |
300 |
apply(case_tac "p = []") |
|
261 | 301 |
apply(simp add: PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
302 |
apply(auto simp add: pflat_len_simps) |
252 | 303 |
apply(rule_tac x="[]" in exI) |
304 |
apply(simp add: Pos_empty pflat_len_simps) |
|
261 | 305 |
apply(auto simp add: pflat_len_simps PosOrd_def) |
252 | 306 |
apply(rule_tac x="ps" in exI) |
307 |
apply(auto) |
|
308 |
apply(drule_tac x="Suc 0#q" in bspec) |
|
309 |
apply(simp) |
|
310 |
apply(simp add: pflat_len_simps) |
|
311 |
apply(drule_tac x="Suc 0#q" in bspec) |
|
312 |
apply(simp) |
|
313 |
apply(simp add: pflat_len_simps) |
|
314 |
done |
|
315 |
||
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
316 |
|
261 | 317 |
lemma PosOrd_SeqI1: |
252 | 318 |
assumes "v1 :\<sqsubset>val v1'" "flat (Seq v1 v2) = flat (Seq v1' v2')" |
319 |
shows "(Seq v1 v2) :\<sqsubset>val (Seq v1' v2')" |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
320 |
using assms(1) |
261 | 321 |
apply(subst (asm) PosOrd_ex_def) |
322 |
apply(subst (asm) PosOrd_def) |
|
252 | 323 |
apply(clarify) |
261 | 324 |
apply(subst PosOrd_ex_def) |
252 | 325 |
apply(rule_tac x="0#p" in exI) |
261 | 326 |
apply(subst PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
327 |
apply(rule conjI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
328 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
329 |
apply(rule conjI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
330 |
apply(simp add: pflat_len_simps) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
331 |
apply(rule ballI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
332 |
apply(rule impI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
333 |
apply(simp only: Pos.simps) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
334 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
335 |
apply(simp add: pflat_len_simps) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
336 |
using assms(2) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
337 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
338 |
apply(auto simp add: pflat_len_simps)[2] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
339 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
340 |
|
261 | 341 |
lemma PosOrd_SeqI2: |
252 | 342 |
assumes "v2 :\<sqsubset>val v2'" "flat v2 = flat v2'" |
343 |
shows "(Seq v v2) :\<sqsubset>val (Seq v v2')" |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
344 |
using assms(1) |
261 | 345 |
apply(subst (asm) PosOrd_ex_def) |
346 |
apply(subst (asm) PosOrd_def) |
|
252 | 347 |
apply(clarify) |
261 | 348 |
apply(subst PosOrd_ex_def) |
252 | 349 |
apply(rule_tac x="Suc 0#p" in exI) |
261 | 350 |
apply(subst PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
351 |
apply(rule conjI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
352 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
353 |
apply(rule conjI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
354 |
apply(simp add: pflat_len_simps) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
355 |
apply(rule ballI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
356 |
apply(rule impI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
357 |
apply(simp only: Pos.simps) |
252 | 358 |
apply(auto)[1] |
359 |
apply(simp add: pflat_len_simps) |
|
360 |
using assms(2) |
|
361 |
apply(simp) |
|
362 |
apply(auto simp add: pflat_len_simps) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
363 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
364 |
|
261 | 365 |
lemma PosOrd_SeqE: |
254 | 366 |
assumes "(Seq v1 v2) :\<sqsubset>val (Seq v1' v2')" |
367 |
shows "v1 :\<sqsubset>val v1' \<or> v2 :\<sqsubset>val v2'" |
|
368 |
using assms |
|
261 | 369 |
apply(simp add: PosOrd_ex_def) |
254 | 370 |
apply(erule exE) |
371 |
apply(case_tac p) |
|
261 | 372 |
apply(simp add: PosOrd_def) |
254 | 373 |
apply(auto simp add: pflat_len_simps intlen_append)[1] |
374 |
apply(rule_tac x="[]" in exI) |
|
375 |
apply(drule_tac x="[]" in spec) |
|
376 |
apply(simp add: Pos_empty pflat_len_simps) |
|
377 |
apply(case_tac a) |
|
378 |
apply(rule disjI1) |
|
261 | 379 |
apply(simp add: PosOrd_def) |
254 | 380 |
apply(auto simp add: pflat_len_simps intlen_append)[1] |
381 |
apply(rule_tac x="list" in exI) |
|
382 |
apply(simp) |
|
383 |
apply(rule ballI) |
|
384 |
apply(rule impI) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
385 |
apply(drule_tac x="0#q" in bspec) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
386 |
apply(simp) |
254 | 387 |
apply(simp add: pflat_len_simps) |
388 |
apply(case_tac nat) |
|
389 |
apply(rule disjI2) |
|
261 | 390 |
apply(simp add: PosOrd_def) |
254 | 391 |
apply(auto simp add: pflat_len_simps intlen_append) |
392 |
apply(rule_tac x="list" in exI) |
|
393 |
apply(simp add: Pos_empty) |
|
394 |
apply(rule ballI) |
|
395 |
apply(rule impI) |
|
396 |
apply(drule_tac x="Suc 0#q" in bspec) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
397 |
apply(simp) |
254 | 398 |
apply(simp add: pflat_len_simps) |
261 | 399 |
apply(simp add: PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
400 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
401 |
|
261 | 402 |
lemma PosOrd_StarsI: |
254 | 403 |
assumes "v1 :\<sqsubset>val v2" "flat (Stars (v1#vs1)) = flat (Stars (v2#vs2))" |
404 |
shows "(Stars (v1#vs1)) :\<sqsubset>val (Stars (v2#vs2))" |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
405 |
using assms(1) |
261 | 406 |
apply(subst (asm) PosOrd_ex_def) |
407 |
apply(subst (asm) PosOrd_def) |
|
254 | 408 |
apply(clarify) |
261 | 409 |
apply(subst PosOrd_ex_def) |
410 |
apply(subst PosOrd_def) |
|
254 | 411 |
apply(rule_tac x="0#p" in exI) |
412 |
apply(simp add: pflat_len_Stars_simps pflat_len_simps) |
|
413 |
using assms(2) |
|
414 |
apply(simp add: pflat_len_simps intlen_append) |
|
415 |
apply(auto simp add: pflat_len_Stars_simps pflat_len_simps) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
416 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
417 |
|
261 | 418 |
lemma PosOrd_StarsI2: |
254 | 419 |
assumes "(Stars vs1) :\<sqsubset>val (Stars vs2)" "flat (Stars vs1) = flat (Stars vs2)" |
420 |
shows "(Stars (v#vs1)) :\<sqsubset>val (Stars (v#vs2))" |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
421 |
using assms(1) |
261 | 422 |
apply(subst (asm) PosOrd_ex_def) |
423 |
apply(subst (asm) PosOrd_def) |
|
254 | 424 |
apply(clarify) |
261 | 425 |
apply(subst PosOrd_ex_def) |
426 |
apply(subst PosOrd_def) |
|
254 | 427 |
apply(case_tac p) |
428 |
apply(simp add: pflat_len_simps) |
|
429 |
apply(rule_tac x="[]" in exI) |
|
430 |
apply(simp add: pflat_len_Stars_simps pflat_len_simps intlen_append) |
|
431 |
apply(rule_tac x="Suc a#list" in exI) |
|
432 |
apply(simp add: pflat_len_Stars_simps pflat_len_simps) |
|
433 |
using assms(2) |
|
434 |
apply(simp add: pflat_len_simps intlen_append) |
|
435 |
apply(auto simp add: pflat_len_Stars_simps pflat_len_simps) |
|
436 |
done |
|
437 |
||
261 | 438 |
lemma PosOrd_Stars_appendI: |
254 | 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) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
443 |
apply(simp) |
261 | 444 |
apply(simp add: PosOrd_StarsI2) |
254 | 445 |
done |
446 |
||
261 | 447 |
lemma PosOrd_StarsE2: |
254 | 448 |
assumes "Stars (v # vs1) :\<sqsubset>val Stars (v # vs2)" |
449 |
shows "Stars vs1 :\<sqsubset>val Stars vs2" |
|
450 |
using assms |
|
261 | 451 |
apply(subst (asm) PosOrd_ex_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
452 |
apply(erule exE) |
254 | 453 |
apply(case_tac p) |
454 |
apply(simp) |
|
261 | 455 |
apply(simp add: PosOrd_def pflat_len_simps intlen_append) |
456 |
apply(subst PosOrd_ex_def) |
|
254 | 457 |
apply(rule_tac x="[]" in exI) |
261 | 458 |
apply(simp add: PosOrd_def pflat_len_simps Pos_empty) |
254 | 459 |
apply(simp) |
460 |
apply(case_tac a) |
|
461 |
apply(clarify) |
|
261 | 462 |
apply(auto simp add: pflat_len_simps PosOrd_def pflat_len_def)[1] |
254 | 463 |
apply(clarify) |
261 | 464 |
apply(simp add: PosOrd_ex_def) |
254 | 465 |
apply(rule_tac x="nat#list" in exI) |
261 | 466 |
apply(auto simp add: PosOrd_def pflat_len_simps intlen_append)[1] |
254 | 467 |
apply(case_tac q) |
261 | 468 |
apply(simp add: PosOrd_def pflat_len_simps intlen_append) |
254 | 469 |
apply(clarify) |
470 |
apply(drule_tac x="Suc a # lista" in bspec) |
|
471 |
apply(simp) |
|
261 | 472 |
apply(auto simp add: PosOrd_def pflat_len_simps intlen_append)[1] |
254 | 473 |
apply(case_tac q) |
261 | 474 |
apply(simp add: PosOrd_def pflat_len_simps intlen_append) |
254 | 475 |
apply(clarify) |
476 |
apply(drule_tac x="Suc a # lista" in bspec) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
477 |
apply(simp) |
261 | 478 |
apply(auto simp add: PosOrd_def pflat_len_simps intlen_append)[1] |
254 | 479 |
done |
480 |
||
261 | 481 |
lemma PosOrd_Stars_appendE: |
254 | 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) |
|
261 | 487 |
apply(simp add: PosOrd_StarsE2) |
254 | 488 |
done |
489 |
||
261 | 490 |
lemma PosOrd_Stars_append_eq: |
254 | 491 |
assumes "flat (Stars vs1) = flat (Stars 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) |
|
261 | 495 |
apply(erule PosOrd_Stars_appendE) |
496 |
apply(rule PosOrd_Stars_appendI) |
|
254 | 497 |
apply(auto) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
498 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
499 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
500 |
|
261 | 501 |
lemma PosOrd_trans: |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
502 |
assumes "v1 :\<sqsubset>val v2" "v2 :\<sqsubset>val v3" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
503 |
shows "v1 :\<sqsubset>val v3" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
504 |
using assms |
261 | 505 |
unfolding PosOrd_ex_def |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
506 |
apply(clarify) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
507 |
apply(subgoal_tac "p = pa \<or> p \<sqsubset>lex pa \<or> pa \<sqsubset>lex p") |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
508 |
prefer 2 |
257 | 509 |
apply(rule lex_trichotomous) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
510 |
apply(erule disjE) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
511 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
512 |
apply(rule_tac x="pa" in exI) |
261 | 513 |
apply(subst PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
514 |
apply(rule conjI) |
261 | 515 |
apply(simp add: PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
516 |
apply(auto)[1] |
261 | 517 |
apply(simp add: PosOrd_def) |
518 |
apply(simp add: PosOrd_def) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
519 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
520 |
using outside_lemma apply blast |
261 | 521 |
apply(simp add: PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
522 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
523 |
using outside_lemma apply force |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
524 |
apply auto[1] |
261 | 525 |
apply(simp add: PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
526 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
527 |
apply (metis (no_types, hide_lams) lex_trans outside_lemma) |
261 | 528 |
apply(simp add: PosOrd_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
529 |
apply(auto)[1] |
255 | 530 |
by (smt intlen_bigger lex_trans outside_lemma pflat_len_def) |
254 | 531 |
|
261 | 532 |
lemma PosOrd_irrefl: |
256 | 533 |
assumes "v :\<sqsubset>val v" |
534 |
shows "False" |
|
535 |
using assms |
|
261 | 536 |
by(auto simp add: PosOrd_ex_def PosOrd_def) |
256 | 537 |
|
261 | 538 |
lemma PosOrd_almost_trichotomous: |
256 | 539 |
shows "v1 :\<sqsubset>val v2 \<or> v2 :\<sqsubset>val v1 \<or> (intlen (flat v1) = intlen (flat v2))" |
261 | 540 |
apply(auto simp add: PosOrd_ex_def) |
541 |
apply(auto simp add: PosOrd_def) |
|
256 | 542 |
apply(rule_tac x="[]" in exI) |
543 |
apply(auto simp add: Pos_empty pflat_len_simps) |
|
544 |
apply(drule_tac x="[]" in spec) |
|
545 |
apply(auto simp add: Pos_empty pflat_len_simps) |
|
546 |
done |
|
547 |
||
548 |
lemma WW1: |
|
549 |
assumes "v1 :\<sqsubset>val v2" "v2 :\<sqsubset>val v1" |
|
550 |
shows "False" |
|
551 |
using assms |
|
261 | 552 |
apply(auto simp add: PosOrd_ex_def PosOrd_def) |
553 |
using assms(1) assms(2) PosOrd_irrefl PosOrd_trans by blast |
|
256 | 554 |
|
555 |
lemma WW2: |
|
556 |
assumes "\<not>(v1 :\<sqsubset>val v2)" |
|
557 |
shows "v1 = v2 \<or> v2 :\<sqsubset>val v1" |
|
558 |
using assms |
|
559 |
oops |
|
560 |
||
261 | 561 |
lemma PosOrd_SeqE2: |
256 | 562 |
assumes "(Seq v1 v2) :\<sqsubset>val (Seq v1' v2')" |
563 |
shows "v1 :\<sqsubset>val v1' \<or> (v1 = v1' \<and> v2 :\<sqsubset>val v2')" |
|
564 |
using assms |
|
261 | 565 |
apply(frule_tac PosOrd_SeqE) |
256 | 566 |
apply(erule disjE) |
567 |
apply(simp) |
|
568 |
apply(auto) |
|
569 |
apply(case_tac "v1 :\<sqsubset>val v1'") |
|
570 |
apply(simp) |
|
261 | 571 |
apply(auto simp add: PosOrd_ex_def) |
256 | 572 |
apply(case_tac "v1 = v1'") |
573 |
apply(simp) |
|
574 |
oops |
|
254 | 575 |
|
576 |
section {* CPT and CPTpre *} |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
577 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
578 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
579 |
inductive |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
580 |
CPrf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<Turnstile> _ : _" [100, 100] 100) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
581 |
where |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
582 |
"\<lbrakk>\<Turnstile> v1 : r1; \<Turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<Turnstile> Seq v1 v2 : SEQ r1 r2" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
583 |
| "\<Turnstile> v1 : r1 \<Longrightarrow> \<Turnstile> Left v1 : ALT r1 r2" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
584 |
| "\<Turnstile> v2 : r2 \<Longrightarrow> \<Turnstile> Right v2 : ALT r1 r2" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
585 |
| "\<Turnstile> Void : ONE" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
586 |
| "\<Turnstile> Char c : CHAR c" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
587 |
| "\<Turnstile> Stars [] : STAR r" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
588 |
| "\<lbrakk>\<Turnstile> v : r; flat v \<noteq> []; \<Turnstile> Stars vs : STAR r\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (v # vs) : STAR r" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
589 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
590 |
lemma Prf_CPrf: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
591 |
assumes "\<Turnstile> v : r" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
592 |
shows "\<turnstile> v : r" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
593 |
using assms |
254 | 594 |
by (induct) (auto intro: Prf.intros) |
595 |
||
256 | 596 |
lemma pflat_len_equal_iff: |
597 |
assumes "\<Turnstile> v1 : r" "\<Turnstile> v2 : r" |
|
598 |
and "\<forall>p. pflat_len v1 p = pflat_len v2 p" |
|
599 |
shows "v1 = v2" |
|
600 |
using assms |
|
601 |
apply(induct v1 r arbitrary: v2 rule: CPrf.induct) |
|
602 |
apply(rotate_tac 4) |
|
603 |
apply(erule CPrf.cases) |
|
604 |
apply(simp_all)[7] |
|
605 |
apply (metis pflat_len_simps(1) pflat_len_simps(2)) |
|
606 |
apply(rotate_tac 2) |
|
607 |
apply(erule CPrf.cases) |
|
608 |
apply(simp_all)[7] |
|
609 |
apply (metis pflat_len_simps(3)) |
|
610 |
apply (metis intlen.simps(1) intlen_length length_greater_0_conv list.size(3) neg_0_le_iff_le not_less not_less_iff_gr_or_eq not_one_le_zero pflat_len_simps(3) pflat_len_simps(6) pflat_len_simps(9)) |
|
611 |
apply(rotate_tac 2) |
|
612 |
apply(erule CPrf.cases) |
|
613 |
apply(simp_all)[7] |
|
614 |
apply (metis intlen.simps(1) intlen_length length_greater_0_conv list.size(3) neg_0_le_iff_le not_less not_less_iff_gr_or_eq not_one_le_zero pflat_len_simps(3) pflat_len_simps(6) pflat_len_simps(9)) |
|
615 |
apply (metis pflat_len_simps(5)) |
|
616 |
apply(erule CPrf.cases) |
|
617 |
apply(simp_all)[7] |
|
618 |
apply(erule CPrf.cases) |
|
619 |
apply(simp_all)[7] |
|
620 |
apply(erule CPrf.cases) |
|
621 |
apply(simp_all)[7] |
|
622 |
apply (metis append_Cons flat.simps(6) flat.simps(7) intlen_length length_greater_0_conv neq_Nil_conv not_less_iff_gr_or_eq pflat_len_simps(9)) |
|
623 |
apply(rotate_tac 5) |
|
624 |
apply(erule CPrf.cases) |
|
625 |
apply(simp_all)[7] |
|
626 |
apply (metis append_Cons flat.simps(6) flat.simps(7) intlen_length length_greater_0_conv list.distinct(1) list.exhaust not_less_iff_gr_or_eq pflat_len_simps(9)) |
|
627 |
apply(auto) |
|
628 |
apply (metis pflat_len_simps(8)) |
|
629 |
apply(subgoal_tac "v = va") |
|
630 |
prefer 2 |
|
631 |
apply (metis pflat_len_simps(8)) |
|
632 |
apply(simp) |
|
633 |
apply(rotate_tac 3) |
|
634 |
apply(drule_tac x="Stars vsa" in meta_spec) |
|
635 |
apply(simp) |
|
636 |
apply(drule_tac meta_mp) |
|
637 |
apply(rule allI) |
|
638 |
apply(case_tac p) |
|
639 |
apply(simp add: pflat_len_simps) |
|
640 |
apply(drule_tac x="[]" in spec) |
|
641 |
apply(simp add: pflat_len_simps intlen_append) |
|
642 |
apply(drule_tac x="Suc a#list" in spec) |
|
643 |
apply(simp add: pflat_len_simps intlen_append) |
|
644 |
apply(simp) |
|
645 |
done |
|
646 |
||
261 | 647 |
lemma PosOrd_trichotomous_stronger: |
256 | 648 |
assumes "\<Turnstile> v1 : r" "\<Turnstile> v2 : r" |
649 |
shows "v1 :\<sqsubset>val v2 \<or> v2 :\<sqsubset>val v1 \<or> (v1 = v2)" |
|
650 |
oops |
|
651 |
||
254 | 652 |
lemma CPrf_stars: |
653 |
assumes "\<Turnstile> Stars vs : STAR r" |
|
654 |
shows "\<forall>v \<in> set vs. flat v \<noteq> [] \<and> \<Turnstile> v : r" |
|
655 |
using assms |
|
656 |
apply(induct vs) |
|
657 |
apply(auto) |
|
658 |
apply(erule CPrf.cases) |
|
659 |
apply(simp_all) |
|
660 |
apply(erule CPrf.cases) |
|
661 |
apply(simp_all) |
|
662 |
apply(erule CPrf.cases) |
|
663 |
apply(simp_all) |
|
664 |
apply(erule CPrf.cases) |
|
665 |
apply(simp_all) |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
666 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
667 |
|
254 | 668 |
lemma CPrf_Stars_appendE: |
669 |
assumes "\<Turnstile> Stars (vs1 @ vs2) : STAR r" |
|
670 |
shows "\<Turnstile> Stars vs1 : STAR r \<and> \<Turnstile> Stars vs2 : STAR r" |
|
671 |
using assms |
|
672 |
apply(induct vs1 arbitrary: vs2) |
|
673 |
apply(auto intro: CPrf.intros)[1] |
|
674 |
apply(erule CPrf.cases) |
|
675 |
apply(simp_all) |
|
676 |
apply(auto intro: CPrf.intros) |
|
677 |
done |
|
678 |
||
679 |
definition PT :: "rexp \<Rightarrow> string \<Rightarrow> val set" |
|
680 |
where "PT r s \<equiv> {v. flat v = s \<and> \<turnstile> v : r}" |
|
681 |
||
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
682 |
definition |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
683 |
"CPT r s = {v. flat v = s \<and> \<Turnstile> v : r}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
684 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
685 |
definition |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
686 |
"CPTpre r s = {v. \<exists>s'. flat v @ s' = s \<and> \<Turnstile> v : r}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
687 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
688 |
lemma CPT_CPTpre_subset: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
689 |
shows "CPT r s \<subseteq> CPTpre r s" |
254 | 690 |
by(auto simp add: CPT_def CPTpre_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
691 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
692 |
lemma CPTpre_subsets: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
693 |
"CPTpre ZERO s = {}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
694 |
"CPTpre ONE s \<subseteq> {Void}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
695 |
"CPTpre (CHAR c) s \<subseteq> {Char c}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
696 |
"CPTpre (ALT r1 r2) s \<subseteq> Left ` CPTpre r1 s \<union> Right ` CPTpre r2 s" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
697 |
"CPTpre (SEQ r1 r2) s \<subseteq> {Seq v1 v2 | v1 v2. v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
698 |
"CPTpre (STAR r) s \<subseteq> {Stars []} \<union> |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
699 |
{Stars (v#vs) | v vs. v \<in> CPTpre r s \<and> flat v \<noteq> [] \<and> Stars vs \<in> CPTpre (STAR r) (drop (length (flat v)) s)}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
700 |
"CPTpre (STAR r) [] = {Stars []}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
701 |
apply(auto simp add: CPTpre_def) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
702 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
703 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
704 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
705 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
706 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
707 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
708 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
709 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
710 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
711 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
712 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
713 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
714 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
715 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
716 |
apply(rule CPrf.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
717 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
718 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
719 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
720 |
lemma CPTpre_simps: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
721 |
shows "CPTpre ONE s = {Void}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
722 |
and "CPTpre (CHAR c) (d#s) = (if c = d then {Char c} else {})" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
723 |
and "CPTpre (ALT r1 r2) s = Left ` CPTpre r1 s \<union> Right ` CPTpre r2 s" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
724 |
and "CPTpre (SEQ r1 r2) s = |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
725 |
{Seq v1 v2 | v1 v2. v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
726 |
apply - |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
727 |
apply(rule subset_antisym) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
728 |
apply(rule CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
729 |
apply(auto simp add: CPTpre_def intro: "CPrf.intros")[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
730 |
apply(case_tac "c = d") |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
731 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
732 |
apply(rule subset_antisym) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
733 |
apply(rule CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
734 |
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
735 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
736 |
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
737 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
738 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
739 |
apply(rule subset_antisym) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
740 |
apply(rule CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
741 |
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
742 |
apply(rule subset_antisym) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
743 |
apply(rule CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
744 |
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
745 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
746 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
747 |
lemma CPT_simps: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
748 |
shows "CPT ONE s = (if s = [] then {Void} else {})" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
749 |
and "CPT (CHAR c) [d] = (if c = d then {Char c} else {})" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
750 |
and "CPT (ALT r1 r2) s = Left ` CPT r1 s \<union> Right ` CPT r2 s" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
751 |
and "CPT (SEQ r1 r2) s = |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
752 |
{Seq v1 v2 | v1 v2 s1 s2. s1 @ s2 = s \<and> v1 \<in> CPT r1 s1 \<and> v2 \<in> CPT r2 s2}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
753 |
apply - |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
754 |
apply(rule subset_antisym) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
755 |
apply(auto simp add: CPT_def)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
756 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
757 |
apply(simp_all)[7] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
758 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
759 |
apply(simp_all)[7] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
760 |
apply(auto simp add: CPT_def intro: CPrf.intros)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
761 |
apply(auto simp add: CPT_def intro: CPrf.intros)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
762 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
763 |
apply(simp_all)[7] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
764 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
765 |
apply(simp_all)[7] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
766 |
apply(auto simp add: CPT_def image_def intro: CPrf.intros)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
767 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
768 |
apply(simp_all)[7] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
769 |
apply(clarify) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
770 |
apply blast |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
771 |
apply(auto simp add: CPT_def image_def intro: CPrf.intros)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
772 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
773 |
apply(simp_all)[7] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
774 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
775 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
776 |
lemma test: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
777 |
assumes "finite A" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
778 |
shows "finite {vs. Stars vs \<in> A}" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
779 |
using assms |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
780 |
apply(induct A) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
781 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
782 |
apply(auto) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
783 |
apply(case_tac x) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
784 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
785 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
786 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
787 |
lemma CPTpre_STAR_finite: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
788 |
assumes "\<And>s. finite (CPTpre r s)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
789 |
shows "finite (CPTpre (STAR r) s)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
790 |
apply(induct s rule: length_induct) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
791 |
apply(case_tac xs) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
792 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
793 |
apply(simp add: CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
794 |
apply(rule finite_subset) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
795 |
apply(rule CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
796 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
797 |
apply(rule_tac B="(\<lambda>(v, vs). Stars (v#vs)) ` {(v, vs). v \<in> CPTpre r (a#list) \<and> flat v \<noteq> [] \<and> Stars vs \<in> CPTpre (STAR r) (drop (length (flat v)) (a#list))}" in finite_subset) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
798 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
799 |
apply(rule finite_imageI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
800 |
apply(simp add: Collect_case_prod_Sigma) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
801 |
apply(rule finite_SigmaI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
802 |
apply(rule assms) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
803 |
apply(case_tac "flat v = []") |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
804 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
805 |
apply(drule_tac x="drop (length (flat v)) (a # list)" in spec) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
806 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
807 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
808 |
apply(rule test) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
809 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
810 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
811 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
812 |
lemma CPTpre_finite: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
813 |
shows "finite (CPTpre r s)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
814 |
apply(induct r arbitrary: s) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
815 |
apply(simp add: CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
816 |
apply(rule finite_subset) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
817 |
apply(rule CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
818 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
819 |
apply(rule finite_subset) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
820 |
apply(rule CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
821 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
822 |
apply(rule finite_subset) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
823 |
apply(rule CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
824 |
apply(rule_tac B="(\<lambda>(v1, v2). Seq v1 v2) ` {(v1, v2). v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}" in finite_subset) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
825 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
826 |
apply(rule finite_imageI) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
827 |
apply(simp add: Collect_case_prod_Sigma) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
828 |
apply(rule finite_subset) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
829 |
apply(rule CPTpre_subsets) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
830 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
831 |
by (simp add: CPTpre_STAR_finite) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
832 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
833 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
834 |
lemma CPT_finite: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
835 |
shows "finite (CPT r s)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
836 |
apply(rule finite_subset) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
837 |
apply(rule CPT_CPTpre_subset) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
838 |
apply(rule CPTpre_finite) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
839 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
840 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
841 |
lemma Posix_CPT: |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
842 |
assumes "s \<in> r \<rightarrow> v" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
843 |
shows "v \<in> CPT r s" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
844 |
using assms |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
845 |
apply(induct rule: Posix.induct) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
846 |
apply(simp add: CPT_def) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
847 |
apply(rule CPrf.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
848 |
apply(simp add: CPT_def) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
849 |
apply(rule CPrf.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
850 |
apply(simp add: CPT_def) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
851 |
apply(rule CPrf.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
852 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
853 |
apply(simp add: CPT_def) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
854 |
apply(rule CPrf.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
855 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
856 |
apply(simp add: CPT_def) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
857 |
apply(rule CPrf.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
858 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
859 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
860 |
apply(simp add: CPT_def) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
861 |
apply(rule CPrf.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
862 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
863 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
864 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
865 |
apply(simp add: CPT_def) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
866 |
apply(rule CPrf.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
867 |
done |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
868 |
|
261 | 869 |
section {* The Posix Value is smaller than any other Value *} |
870 |
||
262 | 871 |
|
261 | 872 |
lemma Posix_PosOrd: |
262 | 873 |
assumes "s \<in> r \<rightarrow> v1" "v2 \<in> CPT r s" |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
874 |
shows "v1 :\<sqsubseteq>val v2" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
875 |
using assms |
261 | 876 |
proof (induct arbitrary: v2 rule: Posix.induct) |
877 |
case (Posix_ONE v) |
|
262 | 878 |
have "v \<in> CPT ONE []" by fact |
879 |
then have "v = Void" |
|
880 |
by (simp add: CPT_simps) |
|
261 | 881 |
then show "Void :\<sqsubseteq>val v" |
262 | 882 |
by (simp add: PosOrd_ex1_def) |
261 | 883 |
next |
884 |
case (Posix_CHAR c v) |
|
262 | 885 |
have "v \<in> CPT (CHAR c) [c]" by fact |
886 |
then have "v = Char c" |
|
887 |
by (simp add: CPT_simps) |
|
261 | 888 |
then show "Char c :\<sqsubseteq>val v" |
262 | 889 |
by (simp add: PosOrd_ex1_def) |
261 | 890 |
next |
891 |
case (Posix_ALT1 s r1 v r2 v2) |
|
892 |
have as1: "s \<in> r1 \<rightarrow> v" by fact |
|
262 | 893 |
have IH: "\<And>v2. v2 \<in> CPT r1 s \<Longrightarrow> v :\<sqsubseteq>val v2" by fact |
894 |
have "v2 \<in> CPT (ALT r1 r2) s" by fact |
|
895 |
then have "\<Turnstile> v2 : ALT r1 r2" "flat v2 = s" |
|
896 |
by(auto simp add: CPT_def prefix_list_def) |
|
261 | 897 |
then consider |
262 | 898 |
(Left) v3 where "v2 = Left v3" "\<Turnstile> v3 : r1" "flat v3 = s" |
899 |
| (Right) v3 where "v2 = Right v3" "\<Turnstile> v3 : r2" "flat v3 = s" |
|
261 | 900 |
by (auto elim: CPrf.cases) |
901 |
then show "Left v :\<sqsubseteq>val v2" |
|
902 |
proof(cases) |
|
903 |
case (Left v3) |
|
262 | 904 |
have "v3 \<in> CPT r1 s" using Left(2,3) |
905 |
by (auto simp add: CPT_def prefix_list_def) |
|
261 | 906 |
with IH have "v :\<sqsubseteq>val v3" by simp |
907 |
moreover |
|
262 | 908 |
have "flat v3 = flat v" using as1 Left(3) |
909 |
by (simp add: Posix1(2)) |
|
261 | 910 |
ultimately have "Left v :\<sqsubseteq>val Left v3" |
262 | 911 |
by (auto simp add: PosOrd_ex1_def PosOrd_LeftI) |
261 | 912 |
then show "Left v :\<sqsubseteq>val v2" unfolding Left . |
913 |
next |
|
914 |
case (Right v3) |
|
262 | 915 |
have "flat v3 = flat v" using as1 Right(3) |
916 |
by (simp add: Posix1(2)) |
|
261 | 917 |
then have "Left v :\<sqsubseteq>val Right v3" using Right(3) as1 |
262 | 918 |
by (auto simp add: PosOrd_ex1_def PosOrd_Left_Right) |
261 | 919 |
then show "Left v :\<sqsubseteq>val v2" unfolding Right . |
920 |
qed |
|
921 |
next |
|
922 |
case (Posix_ALT2 s r2 v r1 v2) |
|
923 |
have as1: "s \<in> r2 \<rightarrow> v" by fact |
|
924 |
have as2: "s \<notin> L r1" by fact |
|
262 | 925 |
have IH: "\<And>v2. v2 \<in> CPT r2 s \<Longrightarrow> v :\<sqsubseteq>val v2" by fact |
926 |
have "v2 \<in> CPT (ALT r1 r2) s" by fact |
|
927 |
then have "\<Turnstile> v2 : ALT r1 r2" "flat v2 = s" |
|
928 |
by(auto simp add: CPT_def prefix_list_def) |
|
261 | 929 |
then consider |
262 | 930 |
(Left) v3 where "v2 = Left v3" "\<Turnstile> v3 : r1" "flat v3 = s" |
931 |
| (Right) v3 where "v2 = Right v3" "\<Turnstile> v3 : r2" "flat v3 = s" |
|
261 | 932 |
by (auto elim: CPrf.cases) |
933 |
then show "Right v :\<sqsubseteq>val v2" |
|
934 |
proof (cases) |
|
935 |
case (Right v3) |
|
262 | 936 |
have "v3 \<in> CPT r2 s" using Right(2,3) |
937 |
by (auto simp add: CPT_def prefix_list_def) |
|
261 | 938 |
with IH have "v :\<sqsubseteq>val v3" by simp |
939 |
moreover |
|
262 | 940 |
have "flat v3 = flat v" using as1 Right(3) |
941 |
by (simp add: Posix1(2)) |
|
261 | 942 |
ultimately have "Right v :\<sqsubseteq>val Right v3" |
262 | 943 |
by (auto simp add: PosOrd_ex1_def PosOrd_RightI) |
261 | 944 |
then show "Right v :\<sqsubseteq>val v2" unfolding Right . |
945 |
next |
|
946 |
case (Left v3) |
|
262 | 947 |
have "v3 \<in> CPT r1 s" using Left(2,3) as2 |
948 |
by (auto simp add: CPT_def prefix_list_def) |
|
949 |
then have "flat v3 = flat v \<and> \<Turnstile> v3 : r1" using as1 Left(3) |
|
950 |
by (simp add: Posix1(2) CPT_def) |
|
951 |
then have "False" using as1 as2 Left |
|
952 |
by (auto simp add: Posix1(2) L_flat_Prf1 Prf_CPrf) |
|
953 |
then show "Right v :\<sqsubseteq>val v2" by simp |
|
261 | 954 |
qed |
955 |
next |
|
956 |
case (Posix_SEQ s1 r1 v1 s2 r2 v2 v3) |
|
957 |
have as1: "s1 \<in> r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" by fact+ |
|
262 | 958 |
have IH1: "\<And>v3. v3 \<in> CPT r1 s1 \<Longrightarrow> v1 :\<sqsubseteq>val v3" by fact |
959 |
have IH2: "\<And>v3. v3 \<in> CPT r2 s2 \<Longrightarrow> v2 :\<sqsubseteq>val v3" by fact |
|
261 | 960 |
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 |
262 | 961 |
have "v3 \<in> CPT (SEQ r1 r2) (s1 @ s2)" by fact |
261 | 962 |
then obtain v3a v3b where eqs: |
963 |
"v3 = Seq v3a v3b" "\<Turnstile> v3a : r1" "\<Turnstile> v3b : r2" |
|
262 | 964 |
"flat v3a @ flat v3b = s1 @ s2" |
965 |
by (force simp add: prefix_list_def CPT_def elim: CPrf.cases) |
|
966 |
with cond have "flat v3a \<sqsubseteq>pre s1" unfolding prefix_list_def |
|
967 |
by (smt L_flat_Prf1 Prf_CPrf append_eq_append_conv2 append_self_conv) |
|
968 |
then have "flat v3a \<sqsubset>spre s1 \<or> (flat v3a = s1 \<and> flat v3b = s2)" using eqs |
|
969 |
by (simp add: sprefix_list_def append_eq_conv_conj) |
|
970 |
then have q2: "v1 :\<sqsubset>val v3a \<or> (flat v3a = s1 \<and> flat v3b = s2)" |
|
971 |
using PosOrd_spreI Posix1(2) as1(1) eqs by blast |
|
972 |
then have "v1 :\<sqsubset>val v3a \<or> (v3a \<in> CPT r1 s1 \<and> v3b \<in> CPT r2 s2)" using eqs(2,3) |
|
973 |
by (auto simp add: CPT_def) |
|
974 |
then have "v1 :\<sqsubset>val v3a \<or> (v1 :\<sqsubseteq>val v3a \<and> v2 :\<sqsubseteq>val v3b)" using IH1 IH2 by blast |
|
975 |
then have "Seq v1 v2 :\<sqsubseteq>val Seq v3a v3b" using eqs q2 as1 |
|
976 |
unfolding PosOrd_ex1_def |
|
977 |
by (metis PosOrd_SeqI1 PosOrd_SeqI2 Posix1(2) flat.simps(5)) |
|
978 |
then show "Seq v1 v2 :\<sqsubseteq>val v3" unfolding eqs by blast |
|
261 | 979 |
next |
980 |
case (Posix_STAR1 s1 r v s2 vs v3) |
|
981 |
have as1: "s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" by fact+ |
|
262 | 982 |
have IH1: "\<And>v3. v3 \<in> CPT r s1 \<Longrightarrow> v :\<sqsubseteq>val v3" by fact |
983 |
have IH2: "\<And>v3. v3 \<in> CPT (STAR r) s2 \<Longrightarrow> Stars vs :\<sqsubseteq>val v3" by fact |
|
261 | 984 |
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 |
985 |
have cond2: "flat v \<noteq> []" by fact |
|
262 | 986 |
have "v3 \<in> CPT (STAR r) (s1 @ s2)" by fact |
261 | 987 |
then consider |
988 |
(NonEmpty) v3a vs3 where |
|
989 |
"v3 = Stars (v3a # vs3)" "\<Turnstile> v3a : r" "\<Turnstile> Stars vs3 : STAR r" |
|
262 | 990 |
"flat (Stars (v3a # vs3)) = s1 @ s2" |
261 | 991 |
| (Empty) "v3 = Stars []" |
262 | 992 |
by (force simp add: CPT_def elim: CPrf.cases) |
261 | 993 |
then show "Stars (v # vs) :\<sqsubseteq>val v3" |
994 |
proof (cases) |
|
995 |
case (NonEmpty v3a vs3) |
|
262 | 996 |
have "flat (Stars (v3a # vs3)) = s1 @ s2" using NonEmpty(4) . |
997 |
with cond have "flat v3a \<sqsubseteq>pre s1" using NonEmpty(2,3) |
|
998 |
unfolding prefix_list_def |
|
999 |
by (smt L_flat_Prf1 Prf_CPrf append_Nil2 append_eq_append_conv2 flat.simps(7)) |
|
1000 |
then have "flat v3a \<sqsubset>spre s1 \<or> (flat v3a = s1 \<and> flat (Stars vs3) = s2)" using NonEmpty(4) |
|
1001 |
by (simp add: sprefix_list_def append_eq_conv_conj) |
|
1002 |
then have q2: "v :\<sqsubset>val v3a \<or> (flat v3a = s1 \<and> flat (Stars vs3) = s2)" |
|
1003 |
using PosOrd_spreI Posix1(2) as1(1) NonEmpty(4) by blast |
|
1004 |
then have "v :\<sqsubset>val v3a \<or> (v3a \<in> CPT r s1 \<and> Stars vs3 \<in> CPT (STAR r) s2)" |
|
1005 |
using NonEmpty(2,3) by (auto simp add: CPT_def) |
|
1006 |
then have "v :\<sqsubset>val v3a \<or> (v :\<sqsubseteq>val v3a \<and> Stars vs :\<sqsubseteq>val Stars vs3)" using IH1 IH2 by blast |
|
1007 |
then have "Stars (v # vs) :\<sqsubseteq>val Stars (v3a # vs3)" using NonEmpty(4) q2 as1 |
|
1008 |
unfolding PosOrd_ex1_def |
|
1009 |
by (metis PosOrd_StarsI PosOrd_StarsI2 Posix1(2) flat.simps(7) val.inject(5)) |
|
1010 |
then show "Stars (v # vs) :\<sqsubseteq>val v3" unfolding NonEmpty by blast |
|
261 | 1011 |
next |
1012 |
case Empty |
|
1013 |
have "v3 = Stars []" by fact |
|
1014 |
then show "Stars (v # vs) :\<sqsubseteq>val v3" |
|
1015 |
unfolding PosOrd_ex1_def using cond2 |
|
1016 |
by (simp add: PosOrd_shorterI) |
|
1017 |
qed |
|
1018 |
next |
|
1019 |
case (Posix_STAR2 r v2) |
|
262 | 1020 |
have "v2 \<in> CPT (STAR r) []" by fact |
1021 |
then have "v2 = Stars []" |
|
1022 |
unfolding CPT_def by (auto elim: CPrf.cases) |
|
261 | 1023 |
then show "Stars [] :\<sqsubseteq>val v2" |
1024 |
by (simp add: PosOrd_ex1_def) |
|
1025 |
qed |
|
253 | 1026 |
|
263 | 1027 |
lemma Posix_PosOrd_stronger: |
1028 |
assumes "s \<in> r \<rightarrow> v1" "v2 \<in> CPTpre r s" |
|
1029 |
shows "v1 :\<sqsubseteq>val v2" |
|
1030 |
proof - |
|
1031 |
from assms(2) have "v2 \<in> CPT r s \<or> flat v2 \<sqsubset>spre s" |
|
1032 |
unfolding CPTpre_def CPT_def sprefix_list_def prefix_list_def by auto |
|
1033 |
moreover |
|
1034 |
{ assume "v2 \<in> CPT r s" |
|
1035 |
with assms(1) have "v1 :\<sqsubseteq>val v2" by (rule Posix_PosOrd) |
|
1036 |
} |
|
1037 |
moreover |
|
1038 |
{ assume "flat v2 \<sqsubset>spre s" |
|
1039 |
then have "flat v2 \<sqsubset>spre flat v1" using assms(1) |
|
1040 |
using Posix1(2) by blast |
|
1041 |
then have "v1 :\<sqsubseteq>val v2" |
|
1042 |
by (simp add: PosOrd_ex1_def PosOrd_spreI) |
|
1043 |
} |
|
1044 |
ultimately show "v1 :\<sqsubseteq>val v2" by blast |
|
1045 |
qed |
|
1046 |
||
261 | 1047 |
lemma Posix_PosOrd_reverse: |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1048 |
assumes "s \<in> r \<rightarrow> v1" |
263 | 1049 |
shows "\<not>(\<exists>v2 \<in> CPTpre r s. v2 :\<sqsubset>val v1)" |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1050 |
using assms |
263 | 1051 |
by (metis Posix_PosOrd_stronger less_irrefl PosOrd_def |
261 | 1052 |
PosOrd_ex1_def PosOrd_ex_def PosOrd_trans) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1053 |
|
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1054 |
|
261 | 1055 |
lemma PosOrd_Posix_Stars: |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1056 |
assumes "(Stars vs) \<in> CPT (STAR r) (flat (Stars vs))" "\<forall>v \<in> set vs. flat v \<in> r \<rightarrow> v" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1057 |
and "\<not>(\<exists>vs2 \<in> PT (STAR r) (flat (Stars vs)). vs2 :\<sqsubset>val (Stars vs))" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1058 |
shows "(flat (Stars vs)) \<in> (STAR r) \<rightarrow> Stars vs" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1059 |
using assms |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1060 |
apply(induct vs) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1061 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1062 |
apply(rule Posix.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1063 |
apply(simp (no_asm)) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1064 |
apply(rule Posix.intros) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1065 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1066 |
apply(auto simp add: CPT_def PT_def)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1067 |
defer |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1068 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1069 |
apply(drule meta_mp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1070 |
apply(auto simp add: CPT_def PT_def)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1071 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1072 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1073 |
apply(drule meta_mp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1074 |
apply(auto simp add: CPT_def PT_def)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1075 |
apply(erule Prf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1076 |
apply(simp_all) |
261 | 1077 |
apply (metis CPrf_stars Cons_eq_map_conv Posix_CPT Posix_STAR2 Posix_PosOrd_reverse list.exhaust list.set_intros(2) map_idI val.distinct(25)) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1078 |
apply(clarify) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1079 |
apply(drule_tac x="Stars (a#v#vsa)" in spec) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1080 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1081 |
apply(drule mp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1082 |
apply (meson CPrf_stars Prf.intros(7) Prf_CPrf list.set_intros(1)) |
261 | 1083 |
apply(subst (asm) (2) PosOrd_ex_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1084 |
apply(simp) |
261 | 1085 |
apply (metis flat.simps(7) flat_Stars PosOrd_StarsI2 PosOrd_ex_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1086 |
apply(auto simp add: CPT_def PT_def)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1087 |
using CPrf_stars apply auto[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1088 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1089 |
apply(auto simp add: CPT_def PT_def)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1090 |
apply(subgoal_tac "\<exists>vA. flat vA = flat a @ s\<^sub>3 \<and> \<turnstile> vA : r") |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1091 |
prefer 2 |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1092 |
apply (meson L_flat_Prf2) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1093 |
apply(subgoal_tac "\<exists>vB. flat (Stars vB) = s\<^sub>4 \<and> \<turnstile> (Stars vB) : (STAR r)") |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1094 |
apply(clarify) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1095 |
apply(drule_tac x="Stars (vA # vB)" in spec) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1096 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1097 |
apply(drule mp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1098 |
using Prf.intros(7) apply blast |
261 | 1099 |
apply(subst (asm) (2) PosOrd_ex_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1100 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1101 |
prefer 2 |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1102 |
apply(simp) |
254 | 1103 |
using Star_values_exists apply blast |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1104 |
prefer 2 |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1105 |
apply(drule meta_mp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1106 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1107 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1108 |
apply(drule meta_mp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1109 |
apply(auto)[1] |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1110 |
prefer 2 |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1111 |
apply(simp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1112 |
apply(erule CPrf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1113 |
apply(simp_all) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1114 |
apply(clarify) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1115 |
apply(rotate_tac 3) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1116 |
apply(erule Prf.cases) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1117 |
apply(simp_all) |
261 | 1118 |
apply (metis CPrf_stars intlen.cases less_irrefl list.set_intros(1) PosOrd_def PosOrd_ex_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1119 |
apply(drule_tac x="Stars (v#va#vsb)" in spec) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1120 |
apply(drule mp) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1121 |
apply (simp add: Posix1a Prf.intros(7)) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1122 |
apply(simp) |
261 | 1123 |
apply(subst (asm) (2) PosOrd_ex_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1124 |
apply(simp) |
261 | 1125 |
apply (metis flat.simps(7) flat_Stars PosOrd_StarsI2 PosOrd_ex_def) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1126 |
proof - |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1127 |
fix a :: val and vsa :: "val list" and s\<^sub>3 :: "char list" and vA :: val and vB :: "val list" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1128 |
assume a1: "s\<^sub>3 \<noteq> []" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1129 |
assume a2: "s\<^sub>3 @ concat (map flat vB) = concat (map flat vsa)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1130 |
assume a3: "flat vA = flat a @ s\<^sub>3" |
261 | 1131 |
assume a4: "\<forall>p. \<not> (Stars (vA # vB) \<sqsubset>val p (Stars (a # vsa)))" |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1132 |
have f5: "\<And>n cs. drop n (cs::char list) = [] \<or> n < length cs" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1133 |
by (meson drop_eq_Nil not_less) |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1134 |
have f6: "\<not> length (flat vA) \<le> length (flat a)" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1135 |
using a3 a1 by simp |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1136 |
have "flat (Stars (a # vsa)) = flat (Stars (vA # vB))" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1137 |
using a3 a2 by simp |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1138 |
then show False |
261 | 1139 |
using f6 f5 a4 by (metis (full_types) drop_eq_Nil PosOrd_StarsI PosOrd_ex_def PosOrd_shorterI) |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1140 |
qed |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1141 |
|
254 | 1142 |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1143 |
|
254 | 1144 |
|
261 | 1145 |
section {* The Smallest Value is indeed the Posix Value *} |
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1146 |
|
261 | 1147 |
lemma PosOrd_Posix: |
1148 |
assumes "v1 \<in> CPT r s" "\<forall>v2 \<in> PT r s. \<not> v2 :\<sqsubset>val v1" |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1149 |
shows "s \<in> r \<rightarrow> v1" |
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1150 |
using assms |
261 | 1151 |
proof(induct r arbitrary: s v1) |
1152 |
case (ZERO s v1) |
|
1153 |
have "v1 \<in> CPT ZERO s" by fact |
|
1154 |
then show "s \<in> ZERO \<rightarrow> v1" unfolding CPT_def |
|
1155 |
by (auto elim: CPrf.cases) |
|
1156 |
next |
|
1157 |
case (ONE s v1) |
|
1158 |
have "v1 \<in> CPT ONE s" by fact |
|
1159 |
then show "s \<in> ONE \<rightarrow> v1" unfolding CPT_def |
|
1160 |
by(auto elim!: CPrf.cases intro: Posix.intros) |
|
1161 |
next |
|
1162 |
case (CHAR c s v1) |
|
1163 |
have "v1 \<in> CPT (CHAR c) s" by fact |
|
1164 |
then show "s \<in> CHAR c \<rightarrow> v1" unfolding CPT_def |
|
1165 |
by (auto elim!: CPrf.cases intro: Posix.intros) |
|
1166 |
next |
|
1167 |
case (ALT r1 r2 s v1) |
|
1168 |
have IH1: "\<And>s v1. \<lbrakk>v1 \<in> CPT r1 s; \<forall>v2 \<in> PT r1 s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r1 \<rightarrow> v1" by fact |
|
1169 |
have IH2: "\<And>s v1. \<lbrakk>v1 \<in> CPT r2 s; \<forall>v2 \<in> PT r2 s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r2 \<rightarrow> v1" by fact |
|
1170 |
have as1: "\<forall>v2\<in>PT (ALT r1 r2) s. \<not> v2 :\<sqsubset>val v1" by fact |
|
1171 |
have as2: "v1 \<in> CPT (ALT r1 r2) s" by fact |
|
1172 |
then consider |
|
1173 |
(Left) v1' where |
|
1174 |
"v1 = Left v1'" "s = flat v1'" |
|
1175 |
"v1' \<in> CPT r1 s" |
|
1176 |
| (Right) v1' where |
|
1177 |
"v1 = Right v1'" "s = flat v1'" |
|
1178 |
"v1' \<in> CPT r2 s" |
|
1179 |
unfolding CPT_def by (auto elim: CPrf.cases) |
|
1180 |
then show "s \<in> ALT r1 r2 \<rightarrow> v1" |
|
1181 |
proof (cases) |
|
1182 |
case (Left v1') |
|
1183 |
have "v1' \<in> CPT r1 s" using as2 |
|
1184 |
unfolding CPT_def Left by (auto elim: CPrf.cases) |
|
1185 |
moreover |
|
1186 |
have "\<forall>v2 \<in> PT r1 s. \<not> v2 :\<sqsubset>val v1'" using as1 |
|
1187 |
unfolding PT_def Left using Prf.intros(2) PosOrd_LeftI by force |
|
1188 |
ultimately have "s \<in> r1 \<rightarrow> v1'" using IH1 by simp |
|
1189 |
then have "s \<in> ALT r1 r2 \<rightarrow> Left v1'" by (rule Posix.intros) |
|
1190 |
then show "s \<in> ALT r1 r2 \<rightarrow> v1" using Left by simp |
|
1191 |
next |
|
1192 |
case (Right v1') |
|
1193 |
have "v1' \<in> CPT r2 s" using as2 |
|
1194 |
unfolding CPT_def Right by (auto elim: CPrf.cases) |
|
1195 |
moreover |
|
1196 |
have "\<forall>v2 \<in> PT r2 s. \<not> v2 :\<sqsubset>val v1'" using as1 |
|
1197 |
unfolding PT_def Right using Prf.intros(3) PosOrd_RightI by force |
|
1198 |
ultimately have "s \<in> r2 \<rightarrow> v1'" using IH2 by simp |
|
1199 |
moreover |
|
1200 |
{ assume "s \<in> L r1" |
|
1201 |
then obtain v' where "v' \<in> PT r1 s" |
|
1202 |
unfolding PT_def using L_flat_Prf2 by blast |
|
1203 |
then have "Left v' \<in> PT (ALT r1 r2) s" |
|
1204 |
unfolding PT_def by (auto intro: Prf.intros) |
|
1205 |
with as1 have "\<not> (Left v' :\<sqsubset>val Right v1') \<and> (flat v' = s)" |
|
1206 |
unfolding PT_def Right by (auto) |
|
1207 |
then have False using PosOrd_Left_Right Right by blast |
|
1208 |
} |
|
1209 |
then have "s \<notin> L r1" by rule |
|
1210 |
ultimately have "s \<in> ALT r1 r2 \<rightarrow> Right v1'" by (rule Posix.intros) |
|
1211 |
then show "s \<in> ALT r1 r2 \<rightarrow> v1" using Right by simp |
|
1212 |
qed |
|
1213 |
next |
|
1214 |
case (SEQ r1 r2 s v1) |
|
1215 |
have IH1: "\<And>s v1. \<lbrakk>v1 \<in> CPT r1 s; \<forall>v2 \<in> PT r1 s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r1 \<rightarrow> v1" by fact |
|
1216 |
have IH2: "\<And>s v1. \<lbrakk>v1 \<in> CPT r2 s; \<forall>v2 \<in> PT r2 s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r2 \<rightarrow> v1" by fact |
|
1217 |
have as1: "\<forall>v2\<in>PT (SEQ r1 r2) s. \<not> v2 :\<sqsubset>val v1" by fact |
|
1218 |
have as2: "v1 \<in> CPT (SEQ r1 r2) s" by fact |
|
1219 |
then obtain |
|
1220 |
v1a v1b where eqs: |
|
1221 |
"v1 = Seq v1a v1b" "s = flat v1a @ flat v1b" |
|
1222 |
"v1a \<in> CPT r1 (flat v1a)" "v1b \<in> CPT r2 (flat v1b)" |
|
1223 |
unfolding CPT_def by(auto elim: CPrf.cases) |
|
1224 |
have "\<forall>v2 \<in> PT r1 (flat v1a). \<not> v2 :\<sqsubset>val v1a" |
|
1225 |
proof |
|
1226 |
fix v2 |
|
1227 |
assume "v2 \<in> PT r1 (flat v1a)" |
|
1228 |
with eqs(2,4) have "Seq v2 v1b \<in> PT (SEQ r1 r2) s" |
|
1229 |
by (simp add: CPT_def PT_def Prf.intros(1) Prf_CPrf) |
|
1230 |
with as1 have "\<not> Seq v2 v1b :\<sqsubset>val Seq v1a v1b \<and> flat (Seq v2 v1b) = flat (Seq v1a v1b)" |
|
1231 |
using eqs by (simp add: PT_def) |
|
1232 |
then show "\<not> v2 :\<sqsubset>val v1a" |
|
1233 |
using PosOrd_SeqI1 by blast |
|
1234 |
qed |
|
1235 |
then have "flat v1a \<in> r1 \<rightarrow> v1a" using IH1 eqs by simp |
|
1236 |
moreover |
|
1237 |
have "\<forall>v2 \<in> PT r2 (flat v1b). \<not> v2 :\<sqsubset>val v1b" |
|
1238 |
proof |
|
1239 |
fix v2 |
|
1240 |
assume "v2 \<in> PT r2 (flat v1b)" |
|
1241 |
with eqs(2,3,4) have "Seq v1a v2 \<in> PT (SEQ r1 r2) s" |
|
1242 |
by (simp add: CPT_def PT_def Prf.intros Prf_CPrf) |
|
1243 |
with as1 have "\<not> Seq v1a v2 :\<sqsubset>val Seq v1a v1b \<and> flat v2 = flat v1b" |
|
1244 |
using eqs by (simp add: PT_def) |
|
1245 |
then show "\<not> v2 :\<sqsubset>val v1b" |
|
1246 |
using PosOrd_SeqI2 by auto |
|
1247 |
qed |
|
1248 |
then have "flat v1b \<in> r2 \<rightarrow> v1b" using IH2 eqs by simp |
|
1249 |
moreover |
|
1250 |
have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = flat v1b \<and> flat v1a @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" |
|
1251 |
proof |
|
1252 |
assume "\<exists>s3 s4. s3 \<noteq> [] \<and> s3 @ s4 = flat v1b \<and> flat v1a @ s3 \<in> L r1 \<and> s4 \<in> L r2" |
|
1253 |
then obtain s3 s4 where q1: "s3 \<noteq> [] \<and> s3 @ s4 = flat v1b \<and> flat v1a @ s3 \<in> L r1 \<and> s4 \<in> L r2" by blast |
|
1254 |
then obtain vA vB where q2: "flat vA = flat v1a @ s3" "\<turnstile> vA : r1" "flat vB = s4" "\<turnstile> vB : r2" |
|
1255 |
using L_flat_Prf2 by blast |
|
1256 |
then have "Seq vA vB \<in> PT (SEQ r1 r2) s" unfolding eqs using q1 |
|
1257 |
by (auto simp add: PT_def intro: Prf.intros) |
|
1258 |
with as1 have "\<not> Seq vA vB :\<sqsubset>val Seq v1a v1b" unfolding eqs by auto |
|
1259 |
then have "\<not> vA :\<sqsubset>val v1a \<and> length (flat vA) > length (flat v1a)" using q1 q2 PosOrd_SeqI1 by auto |
|
1260 |
then show "False" |
|
1261 |
using PosOrd_shorterI by blast |
|
1262 |
qed |
|
1263 |
ultimately |
|
1264 |
show "s \<in> SEQ r1 r2 \<rightarrow> v1" unfolding eqs |
|
1265 |
by (rule Posix.intros) |
|
1266 |
next |
|
1267 |
case (STAR r s v1) |
|
1268 |
have IH: "\<And>s v1. \<lbrakk>v1 \<in> CPT r s; \<forall>v2\<in>PT r s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r \<rightarrow> v1" by fact |
|
1269 |
have as1: "\<forall>v2\<in>PT (STAR r) s. \<not> v2 :\<sqsubset>val v1" by fact |
|
1270 |
have as2: "v1 \<in> CPT (STAR r) s" by fact |
|
1271 |
then obtain |
|
1272 |
vs where eqs: |
|
1273 |
"v1 = Stars vs" "s = flat (Stars vs)" |
|
1274 |
"\<forall>v \<in> set vs. v \<in> CPT r (flat v)" |
|
1275 |
unfolding CPT_def by (auto elim: CPrf.cases dest!: CPrf_stars) |
|
1276 |
have "Stars vs \<in> CPT (STAR r) (flat (Stars vs))" |
|
1277 |
using as2 unfolding eqs . |
|
1278 |
moreover |
|
1279 |
have "\<forall>v\<in>set vs. flat v \<in> r \<rightarrow> v" |
|
1280 |
proof |
|
1281 |
fix v |
|
1282 |
assume a: "v \<in> set vs" |
|
1283 |
then obtain pre post where e: "vs = pre @ [v] @ post" |
|
1284 |
by (metis append_Cons append_Nil in_set_conv_decomp_first) |
|
1285 |
then have q: "\<forall>v2\<in>PT (STAR r) s. \<not> v2 :\<sqsubset>val Stars (pre @ [v] @ post)" |
|
1286 |
using as1 unfolding eqs by simp |
|
1287 |
have "\<forall>v2\<in>PT r (flat v). \<not> v2 :\<sqsubset>val v" unfolding eqs |
|
1288 |
proof (rule ballI, rule notI) |
|
1289 |
fix v2 |
|
1290 |
assume w: "v2 :\<sqsubset>val v" |
|
1291 |
assume "v2 \<in> PT r (flat v)" |
|
1292 |
then have "Stars (pre @ [v2] @ post) \<in> PT (STAR r) s" |
|
1293 |
using as2 unfolding e eqs |
|
1294 |
apply(auto simp add: CPT_def PT_def intro!: Prf_Stars)[1] |
|
1295 |
using CPrf_Stars_appendE CPrf_stars Prf_CPrf apply blast |
|
1296 |
by (meson CPrf_Stars_appendE CPrf_stars Prf_CPrf list.set_intros(2)) |
|
1297 |
then have "\<not> Stars (pre @ [v2] @ post) :\<sqsubset>val Stars (pre @ [v] @ post)" |
|
1298 |
using q by simp |
|
1299 |
with w show "False" |
|
1300 |
using PT_def \<open>v2 \<in> PT r (flat v)\<close> append_Cons flat.simps(7) mem_Collect_eq |
|
1301 |
PosOrd_StarsI PosOrd_Stars_appendI by auto |
|
1302 |
qed |
|
1303 |
with IH |
|
1304 |
show "flat v \<in> r \<rightarrow> v" using a as2 unfolding eqs |
|
1305 |
using eqs(3) by blast |
|
1306 |
qed |
|
1307 |
moreover |
|
1308 |
have "\<not> (\<exists>vs2\<in>PT (STAR r) (flat (Stars vs)). vs2 :\<sqsubset>val Stars vs)" |
|
1309 |
proof |
|
1310 |
assume "\<exists>vs2 \<in> PT (STAR r) (flat (Stars vs)). vs2 :\<sqsubset>val Stars vs" |
|
1311 |
then obtain vs2 where "\<turnstile> Stars vs2 : STAR r" "flat (Stars vs2) = flat (Stars vs)" |
|
1312 |
"Stars vs2 :\<sqsubset>val Stars vs" |
|
1313 |
unfolding PT_def |
|
1314 |
apply(auto elim: Prf.cases) |
|
1315 |
apply(erule Prf.cases) |
|
1316 |
apply(auto intro: Prf.intros) |
|
1317 |
apply(drule_tac x="[]" in meta_spec) |
|
1318 |
apply(simp) |
|
1319 |
apply(drule meta_mp) |
|
1320 |
apply(auto intro: Prf.intros) |
|
1321 |
apply(drule_tac x="(v#vsa)" in meta_spec) |
|
1322 |
apply(auto intro: Prf.intros) |
|
1323 |
done |
|
1324 |
then show "False" using as1 unfolding eqs |
|
1325 |
apply - |
|
1326 |
apply(drule_tac x="Stars vs2" in bspec) |
|
1327 |
apply(auto simp add: PT_def) |
|
1328 |
done |
|
1329 |
qed |
|
1330 |
ultimately have "flat (Stars vs) \<in> STAR r \<rightarrow> Stars vs" |
|
1331 |
by (rule PosOrd_Posix_Stars) |
|
1332 |
then show "s \<in> STAR r \<rightarrow> v1" unfolding eqs . |
|
1333 |
qed |
|
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1334 |
|
249 | 1335 |
unused_thms |
1336 |
||
248
b90ff5abb437
added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1337 |
end |