author | Christian Urban <urbanc@in.tum.de> |
Tue, 27 Jun 2017 13:15:55 +0100 | |
changeset 254 | 7c89d3f6923e |
parent 243 | 09ab631ce7fa |
child 261 | 247fc5dd4943 |
permissions | -rw-r--r-- |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1 |
|
220 | 2 |
theory LexerExt |
185
841f7b9c0a6a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
172
diff
changeset
|
3 |
imports Main |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4 |
begin |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6 |
|
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
7 |
section {* Sequential Composition of Languages *} |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
8 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
9 |
definition |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
10 |
Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100) |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
11 |
where |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
12 |
"A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
13 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
14 |
text {* Two Simple Properties about Sequential Composition *} |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
15 |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
16 |
lemma Sequ_empty [simp]: |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
17 |
shows "A ;; {[]} = A" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
18 |
and "{[]} ;; A = A" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
19 |
by (simp_all add: Sequ_def) |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
20 |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
21 |
lemma Sequ_null [simp]: |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
22 |
shows "A ;; {} = {}" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
23 |
and "{} ;; A = {}" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
24 |
by (simp_all add: Sequ_def) |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
25 |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
26 |
lemma Sequ_assoc: |
220 | 27 |
shows "A ;; (B ;; C) = (A ;; B) ;; C" |
28 |
apply(auto simp add: Sequ_def) |
|
29 |
apply(metis append_assoc) |
|
30 |
apply(metis) |
|
31 |
done |
|
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
32 |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
33 |
lemma Sequ_UNION: |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
34 |
shows "(\<Union>x\<in>A. B ;; C x) = B ;; (\<Union>x\<in>A. C x)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
35 |
by (auto simp add: Sequ_def) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
36 |
|
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
37 |
|
145
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
38 |
section {* Semantic Derivative (Left Quotient) of Languages *} |
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
39 |
|
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
40 |
definition |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
41 |
Der :: "char \<Rightarrow> string set \<Rightarrow> string set" |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
42 |
where |
112
698967eceaf1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
111
diff
changeset
|
43 |
"Der c A \<equiv> {s. c # s \<in> A}" |
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
44 |
|
204
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
45 |
definition |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
46 |
Ders :: "string \<Rightarrow> string set \<Rightarrow> string set" |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
47 |
where |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
48 |
"Ders s A \<equiv> {s'. s @ s' \<in> A}" |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
49 |
|
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
50 |
lemma Der_null [simp]: |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
51 |
shows "Der c {} = {}" |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
52 |
unfolding Der_def |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
53 |
by auto |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
54 |
|
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
55 |
lemma Der_empty [simp]: |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
56 |
shows "Der c {[]} = {}" |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
57 |
unfolding Der_def |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
58 |
by auto |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
59 |
|
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
60 |
lemma Der_char [simp]: |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
61 |
shows "Der c {[d]} = (if c = d then {[]} else {})" |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
62 |
unfolding Der_def |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
63 |
by auto |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
64 |
|
229 | 65 |
lemma Der_Sequ [simp]: |
66 |
shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})" |
|
67 |
unfolding Der_def Sequ_def |
|
68 |
by (auto simp add: Cons_eq_append_conv) |
|
69 |
||
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
70 |
lemma Der_union [simp]: |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
71 |
shows "Der c (A \<union> B) = Der c A \<union> Der c B" |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
72 |
unfolding Der_def |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
73 |
by auto |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
74 |
|
220 | 75 |
lemma Der_UNION: |
76 |
shows "Der c (\<Union>x\<in>A. B x) = (\<Union>x\<in>A. Der c (B x))" |
|
77 |
by (auto simp add: Der_def) |
|
78 |
||
79 |
||
80 |
section {* Power operation for Sets *} |
|
81 |
||
82 |
fun |
|
83 |
Pow :: "string set \<Rightarrow> nat \<Rightarrow> string set" ("_ \<up> _" [101, 102] 101) |
|
84 |
where |
|
85 |
"A \<up> 0 = {[]}" |
|
86 |
| "A \<up> (Suc n) = A ;; (A \<up> n)" |
|
87 |
||
88 |
lemma Pow_empty [simp]: |
|
89 |
shows "[] \<in> A \<up> n \<longleftrightarrow> (n = 0 \<or> [] \<in> A)" |
|
90 |
by(induct n) (auto simp add: Sequ_def) |
|
91 |
||
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
92 |
lemma Der_Pow [simp]: |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
93 |
shows "Der c (A \<up> (Suc n)) = |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
94 |
(Der c A) ;; (A \<up> n) \<union> (if [] \<in> A then Der c (A \<up> n) else {})" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
95 |
unfolding Der_def Sequ_def |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
96 |
by(auto simp add: Cons_eq_append_conv Sequ_def) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
97 |
|
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
98 |
lemma Der_Pow_subseteq: |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
99 |
assumes "[] \<in> A" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
100 |
shows "Der c (A \<up> n) \<subseteq> (Der c A) ;; (A \<up> n)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
101 |
using assms |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
102 |
apply(induct n) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
103 |
apply(simp add: Sequ_def Der_def) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
104 |
apply(simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
105 |
apply(rule conjI) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
106 |
apply (smt Sequ_def append_Nil2 mem_Collect_eq Sequ_assoc subsetI) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
107 |
apply(subgoal_tac "((Der c A) ;; (A \<up> n)) \<subseteq> ((Der c A) ;; (A ;; (A \<up> n)))") |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
108 |
apply(auto)[1] |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
109 |
by (smt Sequ_def append_Nil2 mem_Collect_eq Sequ_assoc subsetI) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
110 |
|
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
111 |
lemma test: |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
112 |
shows "(\<Union>x\<le>Suc n. Der c (A \<up> x)) = (\<Union>x\<le>n. Der c A ;; A \<up> x)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
113 |
apply(induct n) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
114 |
apply(simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
115 |
apply(auto)[1] |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
116 |
apply(case_tac xa) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
117 |
apply(simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
118 |
apply(simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
119 |
apply(auto)[1] |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
120 |
apply(case_tac "[] \<in> A") |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
121 |
apply(simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
122 |
apply(simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
123 |
by (smt Der_Pow Der_Pow_subseteq UN_insert atMost_Suc sup.orderE sup_bot.right_neutral) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
124 |
|
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
125 |
lemma test2: |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
126 |
shows "(\<Union>x\<in>(Suc ` A). Der c (B \<up> x)) = (\<Union>x\<in>A. Der c B ;; B \<up> x)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
127 |
apply(auto)[1] |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
128 |
apply(case_tac "[] \<in> B") |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
129 |
apply(simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
130 |
using Der_Pow_subseteq apply blast |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
131 |
apply(simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
132 |
done |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
133 |
|
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
134 |
|
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
135 |
section {* Kleene Star for Languages *} |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
136 |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
137 |
definition |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
138 |
Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102) where |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
139 |
"A\<star> = (\<Union>n. A \<up> n)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
140 |
|
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
141 |
lemma Star_empty [intro]: |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
142 |
shows "[] \<in> A\<star>" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
143 |
unfolding Star_def |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
144 |
by auto |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
145 |
|
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
146 |
lemma Star_step [intro]: |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
147 |
assumes "s1 \<in> A" "s2 \<in> A\<star>" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
148 |
shows "s1 @ s2 \<in> A\<star>" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
149 |
proof - |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
150 |
from assms obtain n where "s1 \<in>A" "s2 \<in> A \<up> n" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
151 |
unfolding Star_def by auto |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
152 |
then have "s1 @ s2 \<in> A ;; (A \<up> n)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
153 |
by (auto simp add: Sequ_def) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
154 |
then have "s1 @ s2 \<in> A \<up> (Suc n)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
155 |
by simp |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
156 |
then show "s1 @ s2 \<in> A\<star>" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
157 |
unfolding Star_def |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
158 |
by (auto simp del: Pow.simps) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
159 |
qed |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
160 |
|
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
161 |
lemma star_cases: |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
162 |
shows "A\<star> = {[]} \<union> A ;; A\<star>" |
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
163 |
unfolding Star_def |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
164 |
apply(simp add: Sequ_def) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
165 |
apply(auto) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
166 |
apply(case_tac xa) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
167 |
apply(auto simp add: Sequ_def) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
168 |
apply(rule_tac x="Suc xa" in exI) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
169 |
apply(auto simp add: Sequ_def) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
170 |
done |
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
171 |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
172 |
lemma Der_Star1: |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
173 |
shows "Der c (A ;; A\<star>) = (Der c A) ;; A\<star>" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
174 |
proof - |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
175 |
have "(Der c A) ;; A\<star> = (Der c A) ;; (\<Union>n. A \<up> n)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
176 |
unfolding Star_def by simp |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
177 |
also have "... = (\<Union>n. Der c A ;; A \<up> n)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
178 |
unfolding Sequ_UNION by simp |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
179 |
also have "... = (\<Union>x\<in>(Suc ` UNIV). Der c (A \<up> x))" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
180 |
unfolding test2 by simp |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
181 |
also have "... = (\<Union>n. Der c (A \<up> (Suc n)))" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
182 |
by (simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
183 |
also have "... = Der c (\<Union>n. A \<up> (Suc n))" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
184 |
unfolding Der_UNION by simp |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
185 |
also have "... = Der c (A ;; (\<Union>n. A \<up> n))" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
186 |
by (simp only: Pow.simps Sequ_UNION) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
187 |
finally show "Der c (A ;; A\<star>) = (Der c A) ;; A\<star>" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
188 |
unfolding Star_def[symmetric] by blast |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
189 |
qed |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
190 |
|
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
191 |
lemma Der_star [simp]: |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
192 |
shows "Der c (A\<star>) = (Der c A) ;; A\<star>" |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
193 |
proof - |
113
90fe1a1d7d0e
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
112
diff
changeset
|
194 |
have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)" |
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
195 |
by (simp only: star_cases[symmetric]) |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
196 |
also have "... = Der c (A ;; A\<star>)" |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
197 |
by (simp only: Der_union Der_empty) (simp) |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
198 |
also have "... = (Der c A) ;; A\<star>" |
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
199 |
using Der_Star1 by simp |
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
200 |
finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" . |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
201 |
qed |
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
202 |
|
220 | 203 |
|
204 |
||
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
205 |
|
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
206 |
section {* Regular Expressions *} |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
207 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
208 |
datatype rexp = |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
209 |
ZERO |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
210 |
| ONE |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
211 |
| CHAR char |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
212 |
| SEQ rexp rexp |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
213 |
| ALT rexp rexp |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
214 |
| STAR rexp |
220 | 215 |
| UPNTIMES rexp nat |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
216 |
| NTIMES rexp nat |
223 | 217 |
| FROMNTIMES rexp nat |
227 | 218 |
| NMTIMES rexp nat nat |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
219 |
| PLUS rexp |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
220 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
221 |
section {* Semantics of Regular Expressions *} |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
222 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
223 |
fun |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
224 |
L :: "rexp \<Rightarrow> string set" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
225 |
where |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
226 |
"L (ZERO) = {}" |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
227 |
| "L (ONE) = {[]}" |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
228 |
| "L (CHAR c) = {[c]}" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
229 |
| "L (SEQ r1 r2) = (L r1) ;; (L r2)" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
230 |
| "L (ALT r1 r2) = (L r1) \<union> (L r2)" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
231 |
| "L (STAR r) = (L r)\<star>" |
220 | 232 |
| "L (UPNTIMES r n) = (\<Union>i\<in> {..n} . (L r) \<up> i)" |
227 | 233 |
| "L (NTIMES r n) = (L r) \<up> n" |
223 | 234 |
| "L (FROMNTIMES r n) = (\<Union>i\<in> {n..} . (L r) \<up> i)" |
227 | 235 |
| "L (NMTIMES r n m) = (\<Union>i\<in>{n..m} . (L r) \<up> i)" |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
236 |
| "L (PLUS r) = (\<Union>i\<in> {1..} . (L r) \<up> i)" |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
237 |
|
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
238 |
section {* Nullable, Derivatives *} |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
239 |
|
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
240 |
fun |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
241 |
nullable :: "rexp \<Rightarrow> bool" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
242 |
where |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
243 |
"nullable (ZERO) = False" |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
244 |
| "nullable (ONE) = True" |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
245 |
| "nullable (CHAR c) = False" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
246 |
| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
247 |
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)" |
97
38696f516c6b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
248 |
| "nullable (STAR r) = True" |
220 | 249 |
| "nullable (UPNTIMES r n) = True" |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
250 |
| "nullable (NTIMES r n) = (if n = 0 then True else nullable r)" |
223 | 251 |
| "nullable (FROMNTIMES r n) = (if n = 0 then True else nullable r)" |
227 | 252 |
| "nullable (NMTIMES r n m) = (if m < n then False else (if n = 0 then True else nullable r))" |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
253 |
| "nullable (PLUS r) = (nullable r)" |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
254 |
|
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
255 |
fun |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
256 |
der :: "char \<Rightarrow> rexp \<Rightarrow> rexp" |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
257 |
where |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
258 |
"der c (ZERO) = ZERO" |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
259 |
| "der c (ONE) = ZERO" |
111
289728193164
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
110
diff
changeset
|
260 |
| "der c (CHAR d) = (if c = d then ONE else ZERO)" |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
261 |
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)" |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
262 |
| "der c (SEQ r1 r2) = |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
263 |
(if nullable r1 |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
264 |
then ALT (SEQ (der c r1) r2) (der c r2) |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
265 |
else SEQ (der c r1) r2)" |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
266 |
| "der c (STAR r) = SEQ (der c r) (STAR r)" |
220 | 267 |
| "der c (UPNTIMES r 0) = ZERO" |
268 |
| "der c (UPNTIMES r (Suc n)) = SEQ (der c r) (UPNTIMES r n)" |
|
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
269 |
| "der c (NTIMES r 0) = ZERO" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
270 |
| "der c (NTIMES r (Suc n)) = SEQ (der c r) (NTIMES r n)" |
225 | 271 |
| "der c (FROMNTIMES r 0) = SEQ (der c r) (FROMNTIMES r 0)" |
223 | 272 |
| "der c (FROMNTIMES r (Suc n)) = SEQ (der c r) (FROMNTIMES r n)" |
227 | 273 |
| "der c (NMTIMES r n m) = |
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
274 |
(if m < n then ZERO |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
275 |
else (if n = 0 then (if m = 0 then ZERO else |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
276 |
SEQ (der c r) (UPNTIMES r (m - 1))) else |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
277 |
SEQ (der c r) (NMTIMES r (n - 1) (m - 1))))" |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
278 |
| "der c (PLUS r) = SEQ (der c r) (STAR r)" |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
279 |
|
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
280 |
fun |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
281 |
ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp" |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
282 |
where |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
283 |
"ders [] r = r" |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
284 |
| "ders (c # s) r = ders s (der c r)" |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
285 |
|
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
286 |
|
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
287 |
lemma nullable_correctness: |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
288 |
shows "nullable r \<longleftrightarrow> [] \<in> (L r)" |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
289 |
apply(induct r) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
290 |
apply(auto simp add: Sequ_def) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
291 |
done |
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
292 |
|
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
293 |
lemma Der_Pow_notin: |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
294 |
assumes "[] \<notin> A" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
295 |
shows "Der c (A \<up> (Suc n)) = (Der c A) ;; (A \<up> n)" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
296 |
using assms |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
297 |
by(simp) |
220 | 298 |
|
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
299 |
lemma der_correctness: |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
300 |
shows "L (der c r) = Der c (L r)" |
220 | 301 |
apply(induct c r rule: der.induct) |
302 |
apply(simp_all add: nullable_correctness)[7] |
|
303 |
apply(simp only: der.simps L.simps) |
|
304 |
apply(simp only: Der_UNION) |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
305 |
apply(simp only: Sequ_UNION[symmetric]) |
220 | 306 |
apply(simp add: test) |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
307 |
apply(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
308 |
(* NTIMES *) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
309 |
apply(simp only: L.simps der.simps) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
310 |
apply(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
311 |
apply(rule impI) |
223 | 312 |
apply (simp add: Der_Pow_subseteq sup_absorb1) |
313 |
(* FROMNTIMES *) |
|
314 |
apply(simp only: der.simps) |
|
315 |
apply(simp only: L.simps) |
|
225 | 316 |
apply(simp) |
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
317 |
using Der_star Star_def apply auto[1] |
223 | 318 |
apply(simp only: der.simps) |
319 |
apply(simp only: L.simps) |
|
320 |
apply(simp add: Der_UNION) |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
321 |
apply(subst Sequ_UNION[symmetric]) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
322 |
apply(subst test2[symmetric]) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
323 |
apply(subgoal_tac "(Suc ` {n..}) = {Suc n..}") |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
324 |
apply simp |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
325 |
apply(auto simp add: image_def)[1] |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
326 |
using Suc_le_D apply blast |
227 | 327 |
(* NMTIMES *) |
328 |
apply(case_tac "n \<le> m") |
|
329 |
prefer 2 |
|
330 |
apply(simp only: der.simps if_True) |
|
331 |
apply(simp only: L.simps) |
|
332 |
apply(simp) |
|
243 | 333 |
apply(auto) |
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
334 |
apply(subst (asm) Sequ_UNION[symmetric]) |
227 | 335 |
apply(subst (asm) test[symmetric]) |
336 |
apply(auto simp add: Der_UNION)[1] |
|
337 |
apply(auto simp add: Der_UNION)[1] |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
338 |
apply(subst Sequ_UNION[symmetric]) |
227 | 339 |
apply(subst test[symmetric]) |
340 |
apply(auto)[1] |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
341 |
apply(subst (asm) Sequ_UNION[symmetric]) |
227 | 342 |
apply(subst (asm) test2[symmetric]) |
343 |
apply(auto simp add: Der_UNION)[1] |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
344 |
apply(subst Sequ_UNION[symmetric]) |
227 | 345 |
apply(subst test2[symmetric]) |
346 |
apply(auto simp add: Der_UNION)[1] |
|
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
347 |
(* PLUS *) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
348 |
apply(auto simp add: Sequ_def Star_def)[1] |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
349 |
apply(simp add: Der_UNION) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
350 |
apply(rule_tac x="Suc xa" in bexI) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
351 |
apply(auto simp add: Sequ_def Der_def)[2] |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
352 |
apply (metis append_Cons) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
353 |
apply(simp add: Der_UNION) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
354 |
apply(erule_tac bexE) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
355 |
apply(case_tac xa) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
356 |
apply(simp) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
357 |
apply(simp) |
243 | 358 |
apply(auto) |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
359 |
apply(auto simp add: Sequ_def Der_def)[1] |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
360 |
using Star_def apply auto[1] |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
361 |
apply(case_tac "[] \<in> L r") |
243 | 362 |
apply(auto) |
363 |
using Der_UNION Der_star Star_def by fastforce |
|
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
364 |
|
204
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
365 |
lemma ders_correctness: |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
366 |
shows "L (ders s r) = Ders s (L r)" |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
367 |
apply(induct s arbitrary: r) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
368 |
apply(simp_all add: Ders_def der_correctness Der_def) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
369 |
done |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
370 |
|
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
371 |
lemma ders_ZERO: |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
372 |
shows "ders s (ZERO) = ZERO" |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
373 |
apply(induct s) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
374 |
apply(simp_all) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
375 |
done |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
376 |
|
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
377 |
lemma ders_ONE: |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
378 |
shows "ders s (ONE) = (if s = [] then ONE else ZERO)" |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
379 |
apply(induct s) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
380 |
apply(simp_all add: ders_ZERO) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
381 |
done |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
382 |
|
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
383 |
lemma ders_CHAR: |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
384 |
shows "ders s (CHAR c) = (if s = [c] then ONE else |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
385 |
(if s = [] then (CHAR c) else ZERO))" |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
386 |
apply(induct s) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
387 |
apply(simp_all add: ders_ZERO ders_ONE) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
388 |
done |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
389 |
|
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
390 |
lemma ders_ALT: |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
391 |
shows "ders s (ALT r1 r2) = ALT (ders s r1) (ders s r2)" |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
392 |
apply(induct s arbitrary: r1 r2) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
393 |
apply(simp_all) |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
394 |
done |
cd9e40280784
added paper about size derivatives
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
193
diff
changeset
|
395 |
|
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
396 |
section {* Values *} |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
397 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
398 |
datatype val = |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
399 |
Void |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
400 |
| Char char |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
401 |
| Seq val val |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
402 |
| Right val |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
403 |
| Left val |
97
38696f516c6b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
404 |
| Stars "val list" |
243 | 405 |
|
108
73f7dc60c285
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
107
diff
changeset
|
406 |
|
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
407 |
section {* The string behind a value *} |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
408 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
409 |
fun |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
410 |
flat :: "val \<Rightarrow> string" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
411 |
where |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
412 |
"flat (Void) = []" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
413 |
| "flat (Char c) = [c]" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
414 |
| "flat (Left v) = flat v" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
415 |
| "flat (Right v) = flat v" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
416 |
| "flat (Seq v1 v2) = (flat v1) @ (flat v2)" |
97
38696f516c6b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
417 |
| "flat (Stars []) = []" |
38696f516c6b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
418 |
| "flat (Stars (v#vs)) = (flat v) @ (flat (Stars vs))" |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
419 |
|
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
420 |
lemma flat_Stars [simp]: |
97
38696f516c6b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
421 |
"flat (Stars vs) = concat (map flat vs)" |
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
422 |
by (induct vs) (auto) |
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
423 |
|
243 | 424 |
|
425 |
section {* Relation between values and regular expressions *} |
|
426 |
||
427 |
inductive |
|
428 |
Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<turnstile> _ : _" [100, 100] 100) |
|
429 |
where |
|
430 |
"\<lbrakk>\<turnstile> v1 : r1; \<turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<turnstile> Seq v1 v2 : SEQ r1 r2" |
|
431 |
| "\<turnstile> v1 : r1 \<Longrightarrow> \<turnstile> Left v1 : ALT r1 r2" |
|
432 |
| "\<turnstile> v2 : r2 \<Longrightarrow> \<turnstile> Right v2 : ALT r1 r2" |
|
433 |
| "\<turnstile> Void : ONE" |
|
434 |
| "\<turnstile> Char c : CHAR c" |
|
435 |
| "\<lbrakk>\<forall>v \<in> set vs. \<turnstile> v : r\<rbrakk> \<Longrightarrow> \<turnstile> Stars vs : STAR r" |
|
436 |
| "\<lbrakk>\<forall>v \<in> set vs. \<turnstile> v : r; length vs \<le> n\<rbrakk> \<Longrightarrow> \<turnstile> Stars vs : UPNTIMES r n" |
|
437 |
| "\<lbrakk>\<forall>v \<in> set vs. \<turnstile> v : r; length vs = n\<rbrakk> \<Longrightarrow> \<turnstile> Stars vs : NTIMES r n" |
|
438 |
| "\<lbrakk>\<forall>v \<in> set vs. \<turnstile> v : r; length vs \<ge> n\<rbrakk> \<Longrightarrow> \<turnstile> Stars vs : FROMNTIMES r n" |
|
439 |
| "\<lbrakk>\<forall>v \<in> set vs. \<turnstile> v : r; length vs \<ge> n; length vs \<le> m\<rbrakk> \<Longrightarrow> \<turnstile> Stars vs : NMTIMES r n m" |
|
440 |
| "\<lbrakk>\<forall>v \<in> set vs. \<turnstile> v : r; length vs \<ge> 1\<rbrakk> \<Longrightarrow> \<turnstile> Stars vs : PLUS r" |
|
441 |
||
442 |
||
443 |
inductive_cases Prf_elims: |
|
444 |
"\<turnstile> v : ZERO" |
|
445 |
"\<turnstile> v : SEQ r1 r2" |
|
446 |
"\<turnstile> v : ALT r1 r2" |
|
447 |
"\<turnstile> v : ONE" |
|
448 |
"\<turnstile> v : CHAR c" |
|
449 |
(* "\<turnstile> vs : STAR r"*) |
|
450 |
||
451 |
lemma Prf_STAR: |
|
452 |
assumes "\<forall>v\<in>set vs. \<turnstile> v : r \<and> flat v \<in> L r" |
|
453 |
shows "concat (map flat vs) \<in> L r\<star>" |
|
454 |
using assms |
|
455 |
apply(induct vs) |
|
456 |
apply(auto) |
|
457 |
done |
|
458 |
||
459 |
lemma Prf_Pow: |
|
460 |
assumes "\<forall>v\<in>set vs. \<turnstile> v : r \<and> flat v \<in> L r" |
|
461 |
shows "concat (map flat vs) \<in> L r \<up> length vs" |
|
462 |
using assms |
|
463 |
apply(induct vs) |
|
464 |
apply(auto simp add: Sequ_def) |
|
465 |
done |
|
466 |
||
467 |
lemma Prf_flat_L: |
|
468 |
assumes "\<turnstile> v : r" shows "flat v \<in> L r" |
|
469 |
using assms |
|
470 |
apply(induct v r rule: Prf.induct) |
|
471 |
apply(auto simp add: Sequ_def) |
|
472 |
apply(rule Prf_STAR) |
|
473 |
apply(simp) |
|
474 |
apply(rule_tac x="length vs" in bexI) |
|
475 |
apply(rule Prf_Pow) |
|
476 |
apply(simp) |
|
477 |
apply(simp) |
|
478 |
apply(rule Prf_Pow) |
|
479 |
apply(simp) |
|
480 |
apply(rule_tac x="length vs" in bexI) |
|
481 |
apply(rule Prf_Pow) |
|
482 |
apply(simp) |
|
483 |
apply(simp) |
|
484 |
apply(rule_tac x="length vs" in bexI) |
|
485 |
apply(rule Prf_Pow) |
|
486 |
apply(simp) |
|
487 |
apply(simp) |
|
488 |
using Prf_Pow by blast |
|
489 |
||
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
490 |
lemma Star_Pow: |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
491 |
assumes "s \<in> A \<up> n" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
492 |
shows "\<exists>ss. concat ss = s \<and> (\<forall>s \<in> set ss. s \<in> A) \<and> (length ss = n)" |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
493 |
using assms |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
494 |
apply(induct n arbitrary: s) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
495 |
apply(auto simp add: Sequ_def) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
496 |
apply(drule_tac x="s2" in meta_spec) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
497 |
apply(auto) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
498 |
apply(rule_tac x="s1#ss" in exI) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
499 |
apply(simp) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
500 |
done |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
501 |
|
90
3c8cfdf95252
proved some lemmas about star and mkeps (injval etc not yet done)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
89
diff
changeset
|
502 |
lemma Star_string: |
3c8cfdf95252
proved some lemmas about star and mkeps (injval etc not yet done)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
89
diff
changeset
|
503 |
assumes "s \<in> A\<star>" |
3c8cfdf95252
proved some lemmas about star and mkeps (injval etc not yet done)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
89
diff
changeset
|
504 |
shows "\<exists>ss. concat ss = s \<and> (\<forall>s \<in> set ss. s \<in> A)" |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
505 |
using assms |
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
506 |
apply(auto simp add: Star_def) |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
507 |
using Star_Pow by blast |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
508 |
|
90
3c8cfdf95252
proved some lemmas about star and mkeps (injval etc not yet done)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
89
diff
changeset
|
509 |
|
243 | 510 |
lemma Star_val: |
511 |
assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<turnstile> v : r" |
|
512 |
shows "\<exists>vs. concat (map flat vs) = concat ss \<and> (\<forall>v\<in>set vs. \<turnstile> v : r)" |
|
513 |
using assms |
|
514 |
apply(induct ss) |
|
515 |
apply(auto) |
|
516 |
apply (metis empty_iff list.set(1)) |
|
517 |
by (metis concat.simps(2) list.simps(9) set_ConsD) |
|
518 |
||
519 |
lemma Star_val_length: |
|
520 |
assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<turnstile> v : r" |
|
521 |
shows "\<exists>vs. concat (map flat vs) = concat ss \<and> (\<forall>v\<in>set vs. \<turnstile> v : r) \<and> (length vs) = (length ss)" |
|
522 |
using assms |
|
523 |
apply(induct ss) |
|
524 |
apply(auto) |
|
525 |
by (metis List.bind_def bind_simps(2) length_Suc_conv set_ConsD) |
|
526 |
||
527 |
||
528 |
||
529 |
||
530 |
||
531 |
lemma L_flat_Prf2: |
|
532 |
assumes "s \<in> L r" shows "\<exists>v. \<turnstile> v : r \<and> flat v = s" |
|
533 |
using assms |
|
534 |
apply(induct r arbitrary: s) |
|
535 |
apply(auto simp add: Sequ_def intro: Prf.intros) |
|
536 |
using Prf.intros(1) flat.simps(5) apply blast |
|
537 |
using Prf.intros(2) flat.simps(3) apply blast |
|
538 |
using Prf.intros(3) flat.simps(4) apply blast |
|
539 |
apply(subgoal_tac "\<exists>vs::val list. concat (map flat vs) = s \<and> (\<forall>v \<in> set vs. \<turnstile> v : r)") |
|
540 |
apply(auto)[1] |
|
541 |
apply(rule_tac x="Stars vs" in exI) |
|
542 |
apply(simp) |
|
543 |
apply(rule Prf.intros) |
|
544 |
apply(simp) |
|
545 |
using Star_string Star_val apply force |
|
546 |
apply(subgoal_tac "\<exists>vs::val list. concat (map flat vs) = s \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> (length vs = x)") |
|
547 |
apply(auto)[1] |
|
548 |
apply(rule_tac x="Stars vs" in exI) |
|
549 |
apply(simp) |
|
550 |
apply(rule Prf.intros) |
|
551 |
apply(simp) |
|
552 |
apply(simp) |
|
553 |
using Star_Pow Star_val_length apply blast |
|
554 |
apply(subgoal_tac "\<exists>vs::val list. concat (map flat vs) = s \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> (length vs = x2)") |
|
555 |
apply(auto)[1] |
|
556 |
apply(rule_tac x="Stars vs" in exI) |
|
557 |
apply(simp) |
|
558 |
apply(rule Prf.intros) |
|
559 |
apply(simp) |
|
560 |
apply(simp) |
|
561 |
using Star_Pow Star_val_length apply blast |
|
562 |
apply(subgoal_tac "\<exists>vs::val list. concat (map flat vs) = s \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> (length vs = x)") |
|
563 |
apply(auto)[1] |
|
564 |
apply(rule_tac x="Stars vs" in exI) |
|
565 |
apply(simp) |
|
566 |
apply(rule Prf.intros) |
|
567 |
apply(simp) |
|
568 |
apply(simp) |
|
569 |
using Star_Pow Star_val_length apply blast |
|
570 |
apply(subgoal_tac "\<exists>vs::val list. concat (map flat vs) = s \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> (length vs = x)") |
|
571 |
apply(auto)[1] |
|
572 |
apply(rule_tac x="Stars vs" in exI) |
|
573 |
apply(simp) |
|
574 |
apply(rule Prf.intros) |
|
575 |
apply(simp) |
|
576 |
apply(simp) |
|
577 |
apply(simp) |
|
578 |
using Star_Pow Star_val_length apply blast |
|
579 |
apply(subgoal_tac "\<exists>vs::val list. concat (map flat vs) = s \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> (length vs = x)") |
|
580 |
apply(auto)[1] |
|
581 |
apply(rule_tac x="Stars vs" in exI) |
|
582 |
apply(simp) |
|
583 |
apply(rule Prf.intros) |
|
584 |
apply(simp) |
|
585 |
apply(simp) |
|
586 |
using Star_Pow Star_val_length apply blast |
|
587 |
done |
|
588 |
||
589 |
lemma L_flat_Prf: |
|
590 |
"L(r) = {flat v | v. \<turnstile> v : r}" |
|
591 |
using Prf_flat_L L_flat_Prf2 by blast |
|
592 |
||
593 |
||
145
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
594 |
section {* Sulzmann and Lu functions *} |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
595 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
596 |
fun |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
597 |
mkeps :: "rexp \<Rightarrow> val" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
598 |
where |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
599 |
"mkeps(ONE) = Void" |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
600 |
| "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)" |
227 | 601 |
| "mkeps(ALT r1 r2) = (if nullable r1 then Left (mkeps r1) else Right (mkeps r2))" |
97
38696f516c6b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
602 |
| "mkeps(STAR r) = Stars []" |
220 | 603 |
| "mkeps(UPNTIMES r n) = Stars []" |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
604 |
| "mkeps(NTIMES r n) = Stars (replicate n (mkeps r))" |
223 | 605 |
| "mkeps(FROMNTIMES r n) = Stars (replicate n (mkeps r))" |
227 | 606 |
| "mkeps(NMTIMES r n m) = Stars (replicate n (mkeps r))" |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
607 |
| "mkeps(PLUS r) = Stars ([mkeps r])" |
227 | 608 |
|
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
609 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
610 |
fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
611 |
where |
101
7f4f8c34da95
fixed inj function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
100
diff
changeset
|
612 |
"injval (CHAR d) c Void = Char d" |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
613 |
| "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
614 |
| "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
615 |
| "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
616 |
| "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
617 |
| "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)" |
97
38696f516c6b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
618 |
| "injval (STAR r) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" |
220 | 619 |
| "injval (UPNTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
620 |
| "injval (NTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" |
223 | 621 |
| "injval (FROMNTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" |
227 | 622 |
| "injval (NMTIMES r n m) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
623 |
| "injval (PLUS r) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" |
239
e59cec211f4f
added automata implementation
Christian Urban <urbanc@in.tum.de>
parents:
238
diff
changeset
|
624 |
|
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
625 |
section {* Mkeps, injval *} |
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
626 |
|
243 | 627 |
lemma mkeps_nullable: |
628 |
assumes "nullable(r)" |
|
629 |
shows "\<turnstile> mkeps r : r" |
|
630 |
using assms |
|
631 |
apply(induct r rule: mkeps.induct) |
|
632 |
apply(auto intro: Prf.intros) |
|
633 |
by (metis Prf.intros(10) in_set_replicate length_replicate not_le order_refl) |
|
634 |
||
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
635 |
lemma mkeps_flat: |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
636 |
assumes "nullable(r)" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
637 |
shows "flat (mkeps r) = []" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
638 |
using assms |
243 | 639 |
apply (induct rule: nullable.induct) |
640 |
apply(auto) |
|
641 |
by meson |
|
238
2dc1647eab9e
added AND-regular expression (intersection/conjunction)
Christian Urban <urbanc@in.tum.de>
parents:
236
diff
changeset
|
642 |
|
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
643 |
|
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
644 |
lemma Prf_injval: |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
645 |
assumes "\<turnstile> v : der c r" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
646 |
shows "\<turnstile> (injval r c v) : r" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
647 |
using assms |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
648 |
apply(induct r arbitrary: c v rule: rexp.induct) |
142
08dcf0d20f15
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
649 |
apply(auto intro!: Prf.intros mkeps_nullable elim!: Prf_elims split: if_splits) |
107
6adda4a667b1
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
106
diff
changeset
|
650 |
(* STAR *) |
91
f067e59b58d9
more lemmas for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
90
diff
changeset
|
651 |
apply(rotate_tac 2) |
f067e59b58d9
more lemmas for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
90
diff
changeset
|
652 |
apply(erule Prf.cases) |
223 | 653 |
apply(simp_all) |
91
f067e59b58d9
more lemmas for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
90
diff
changeset
|
654 |
apply(auto) |
223 | 655 |
using Prf.intros(6) apply auto[1] |
220 | 656 |
(* UPNTIMES *) |
657 |
apply(case_tac x2) |
|
658 |
apply(simp) |
|
659 |
using Prf_elims(1) apply auto[1] |
|
660 |
apply(simp) |
|
661 |
apply(erule Prf.cases) |
|
662 |
apply(simp_all) |
|
663 |
apply(clarify) |
|
223 | 664 |
apply(rotate_tac 2) |
665 |
apply(erule Prf.cases) |
|
666 |
apply(simp_all) |
|
220 | 667 |
apply(clarify) |
223 | 668 |
using Prf.intros(7) apply auto[1] |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
669 |
(* NTIMES *) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
670 |
apply(case_tac x2) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
671 |
apply(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
672 |
using Prf_elims(1) apply auto[1] |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
673 |
apply(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
674 |
apply(erule Prf.cases) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
675 |
apply(simp_all) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
676 |
apply(clarify) |
223 | 677 |
apply(rotate_tac 2) |
678 |
apply(erule Prf.cases) |
|
679 |
apply(simp_all) |
|
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
680 |
apply(clarify) |
223 | 681 |
using Prf.intros(8) apply auto[1] |
682 |
(* FROMNTIMES *) |
|
683 |
apply(case_tac x2) |
|
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
684 |
apply(simp) |
223 | 685 |
apply(erule Prf.cases) |
686 |
apply(simp_all) |
|
687 |
apply(clarify) |
|
688 |
apply(rotate_tac 2) |
|
689 |
apply(erule Prf.cases) |
|
690 |
apply(simp_all) |
|
691 |
apply(clarify) |
|
692 |
using Prf.intros(9) apply auto[1] |
|
693 |
apply(rotate_tac 1) |
|
694 |
apply(erule Prf.cases) |
|
695 |
apply(simp_all) |
|
696 |
apply(clarify) |
|
697 |
apply(rotate_tac 2) |
|
698 |
apply(erule Prf.cases) |
|
699 |
apply(simp_all) |
|
700 |
apply(clarify) |
|
227 | 701 |
using Prf.intros(9) apply auto |
702 |
(* NMTIMES *) |
|
703 |
apply(rotate_tac 3) |
|
704 |
apply(erule Prf.cases) |
|
705 |
apply(simp_all) |
|
706 |
apply(clarify) |
|
707 |
apply(rule Prf.intros) |
|
708 |
apply(auto)[2] |
|
709 |
apply simp |
|
710 |
apply(rotate_tac 4) |
|
711 |
apply(erule Prf.cases) |
|
712 |
apply(simp_all) |
|
713 |
apply(clarify) |
|
714 |
apply(rule Prf.intros) |
|
715 |
apply(auto)[2] |
|
716 |
apply simp |
|
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
717 |
(* PLUS *) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
718 |
apply(rotate_tac 2) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
719 |
apply(erule Prf.cases) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
720 |
apply(simp_all) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
721 |
apply(auto) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
722 |
using Prf.intros apply auto[1] |
227 | 723 |
done |
223 | 724 |
|
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
725 |
|
243 | 726 |
lemma Prf_injval_flat: |
727 |
assumes "\<turnstile> v : der c r" |
|
728 |
shows "flat (injval r c v) = c # (flat v)" |
|
729 |
using assms |
|
730 |
apply(induct arbitrary: v rule: der.induct) |
|
731 |
apply(auto elim!: Prf_elims split: if_splits) |
|
732 |
apply(metis mkeps_flat) |
|
733 |
apply(rotate_tac 2) |
|
734 |
apply(erule Prf.cases) |
|
735 |
apply(simp_all) |
|
736 |
apply(rotate_tac 2) |
|
737 |
apply(erule Prf.cases) |
|
738 |
apply(simp_all) |
|
739 |
apply(rotate_tac 2) |
|
740 |
apply(erule Prf.cases) |
|
741 |
apply(simp_all) |
|
742 |
apply(rotate_tac 2) |
|
743 |
apply(erule Prf.cases) |
|
744 |
apply(simp_all) |
|
745 |
apply(rotate_tac 2) |
|
746 |
apply(erule Prf.cases) |
|
747 |
apply(simp_all) |
|
748 |
apply(rotate_tac 3) |
|
749 |
apply(erule Prf.cases) |
|
750 |
apply(simp_all) |
|
751 |
apply(rotate_tac 4) |
|
752 |
apply(erule Prf.cases) |
|
753 |
apply(simp_all) |
|
754 |
apply(rotate_tac 2) |
|
755 |
apply(erule Prf.cases) |
|
756 |
apply(simp_all) |
|
757 |
done |
|
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
758 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
759 |
|
104
59bad592a009
updated theories and cleaned them up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
103
diff
changeset
|
760 |
section {* Our Alternative Posix definition *} |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
761 |
|
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
762 |
inductive |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
763 |
Posix :: "string \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<in> _ \<rightarrow> _" [100, 100, 100] 100) |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
764 |
where |
151
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
765 |
Posix_ONE: "[] \<in> ONE \<rightarrow> Void" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
766 |
| Posix_CHAR: "[c] \<in> (CHAR c) \<rightarrow> (Char c)" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
767 |
| Posix_ALT1: "s \<in> r1 \<rightarrow> v \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Left v)" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
768 |
| Posix_ALT2: "\<lbrakk>s \<in> r2 \<rightarrow> v; s \<notin> L(r1)\<rbrakk> \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Right v)" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
769 |
| Posix_SEQ: "\<lbrakk>s1 \<in> r1 \<rightarrow> v1; s2 \<in> r2 \<rightarrow> v2; |
97
38696f516c6b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
770 |
\<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)\<rbrakk> \<Longrightarrow> |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
771 |
(s1 @ s2) \<in> (SEQ r1 r2) \<rightarrow> (Seq v1 v2)" |
151
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
772 |
| Posix_STAR1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> STAR r \<rightarrow> Stars vs; flat v \<noteq> []; |
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
773 |
\<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))\<rbrakk> |
97
38696f516c6b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
774 |
\<Longrightarrow> (s1 @ s2) \<in> STAR r \<rightarrow> Stars (v # vs)" |
151
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
775 |
| Posix_STAR2: "[] \<in> STAR r \<rightarrow> Stars []" |
220 | 776 |
| Posix_UPNTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> UPNTIMES r n \<rightarrow> Stars vs; flat v \<noteq> []; |
777 |
\<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 (UPNTIMES r n))\<rbrakk> |
|
778 |
\<Longrightarrow> (s1 @ s2) \<in> UPNTIMES r (Suc n) \<rightarrow> Stars (v # vs)" |
|
779 |
| Posix_UPNTIMES2: "[] \<in> UPNTIMES r n \<rightarrow> Stars []" |
|
236 | 780 |
| Posix_NTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> NTIMES r n \<rightarrow> Stars vs; flat v = [] \<Longrightarrow> flat (Stars vs) = []; |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
781 |
\<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 (NTIMES r n))\<rbrakk> |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
782 |
\<Longrightarrow> (s1 @ s2) \<in> NTIMES r (Suc n) \<rightarrow> Stars (v # vs)" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
783 |
| Posix_NTIMES2: "[] \<in> NTIMES r 0 \<rightarrow> Stars []" |
236 | 784 |
| Posix_FROMNTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> FROMNTIMES r n \<rightarrow> Stars vs; flat v = [] \<Longrightarrow> flat (Stars vs) = []; |
223 | 785 |
\<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 (FROMNTIMES r n))\<rbrakk> |
786 |
\<Longrightarrow> (s1 @ s2) \<in> FROMNTIMES r (Suc n) \<rightarrow> Stars (v # vs)" |
|
225 | 787 |
| Posix_FROMNTIMES2: "s \<in> STAR r \<rightarrow> Stars vs \<Longrightarrow> s \<in> FROMNTIMES r 0 \<rightarrow> Stars vs" |
234 | 788 |
| Posix_NMTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> NMTIMES r n m \<rightarrow> Stars vs; flat v = [] \<Longrightarrow> flat (Stars vs) = []; |
227 | 789 |
\<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 (NMTIMES r n m))\<rbrakk> |
790 |
\<Longrightarrow> (s1 @ s2) \<in> NMTIMES r (Suc n) (Suc m) \<rightarrow> Stars (v # vs)" |
|
791 |
| Posix_NMTIMES2: "s \<in> UPNTIMES r m \<rightarrow> Stars vs \<Longrightarrow> s \<in> NMTIMES r 0 m \<rightarrow> Stars vs" |
|
233
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
792 |
| Posix_PLUS1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> STAR r \<rightarrow> Stars vs; flat v = [] \<Longrightarrow> flat (Stars vs) = []; |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
793 |
\<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))\<rbrakk> |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
794 |
\<Longrightarrow> (s1 @ s2) \<in> PLUS r \<rightarrow> Stars (v # vs)" |
225 | 795 |
|
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
796 |
inductive_cases Posix_elims: |
149
ec3d221bfc45
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
146
diff
changeset
|
797 |
"s \<in> ZERO \<rightarrow> v" |
143
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
798 |
"s \<in> ONE \<rightarrow> v" |
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
799 |
"s \<in> CHAR c \<rightarrow> v" |
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
800 |
"s \<in> ALT r1 r2 \<rightarrow> v" |
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
801 |
"s \<in> SEQ r1 r2 \<rightarrow> v" |
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
802 |
"s \<in> STAR r \<rightarrow> v" |
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
803 |
|
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
804 |
lemma Posix1: |
101
7f4f8c34da95
fixed inj function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
100
diff
changeset
|
805 |
assumes "s \<in> r \<rightarrow> v" |
123
1bee7006b92b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
122
diff
changeset
|
806 |
shows "s \<in> L r" "flat v = s" |
101
7f4f8c34da95
fixed inj function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
100
diff
changeset
|
807 |
using assms |
220 | 808 |
apply (induct s r v rule: Posix.induct) |
809 |
apply(auto simp add: Sequ_def) |
|
810 |
apply(rule_tac x="Suc x" in bexI) |
|
811 |
apply(auto simp add: Sequ_def) |
|
223 | 812 |
apply(rule_tac x="Suc x" in bexI) |
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
813 |
using Sequ_def apply auto[2] |
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
814 |
apply(simp add: Star_def) |
227 | 815 |
apply(rule_tac x="Suc x" in bexI) |
816 |
apply(auto simp add: Sequ_def) |
|
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
817 |
apply(simp add: Star_def) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
818 |
apply(clarify) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
819 |
apply(rule_tac x="Suc x" in bexI) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
820 |
apply(auto simp add: Sequ_def) |
227 | 821 |
done |
143
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
822 |
|
101
7f4f8c34da95
fixed inj function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
100
diff
changeset
|
823 |
|
233
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
824 |
lemma |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
825 |
"([] @ []) \<in> PLUS (ONE) \<rightarrow> Stars ([Void])" |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
826 |
apply(rule Posix_PLUS1) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
827 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
828 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
829 |
apply(simp) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
830 |
apply(simp) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
831 |
done |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
832 |
|
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
833 |
lemma |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
834 |
assumes "s \<in> r \<rightarrow> v" "flat v \<noteq> []" "\<forall>s' \<in> L r. length s' < length s" |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
835 |
shows "([] @ (s @ [])) \<in> PLUS (ALT ONE r) \<rightarrow> Stars ([Left Void, Right v])" |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
836 |
using assms |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
837 |
oops |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
838 |
|
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
839 |
lemma |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
840 |
"([] @ ([] @ [])) \<in> FROMNTIMES (ONE) (Suc (Suc 0)) \<rightarrow> Stars ([Void, Void])" |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
841 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
842 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
843 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
844 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
845 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
846 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
847 |
apply(auto) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
848 |
done |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
849 |
|
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
850 |
|
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
851 |
lemma |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
852 |
assumes "s \<in> r \<rightarrow> v" "flat v \<noteq> []" |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
853 |
"s \<in> PLUS (ALT ONE r) \<rightarrow> Stars ([Left(Void), Right(v)])" |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
854 |
shows "False" |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
855 |
using assms |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
856 |
apply(rotate_tac 2) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
857 |
apply(erule_tac Posix.cases) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
858 |
apply(simp_all) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
859 |
apply(auto) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
860 |
oops |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
861 |
|
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
862 |
|
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
863 |
|
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
864 |
|
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
865 |
|
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
866 |
lemma Posix1a: |
122
7c6c907660d8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
121
diff
changeset
|
867 |
assumes "s \<in> r \<rightarrow> v" |
123
1bee7006b92b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
122
diff
changeset
|
868 |
shows "\<turnstile> v : r" |
122
7c6c907660d8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
121
diff
changeset
|
869 |
using assms |
220 | 870 |
apply(induct s r v rule: Posix.induct) |
871 |
apply(auto intro: Prf.intros) |
|
223 | 872 |
apply(rule Prf.intros) |
873 |
apply(auto)[1] |
|
874 |
apply(rotate_tac 3) |
|
875 |
apply(erule Prf.cases) |
|
876 |
apply(simp_all) |
|
877 |
apply(rule Prf.intros) |
|
878 |
apply(auto)[1] |
|
879 |
apply(rotate_tac 3) |
|
880 |
apply(erule Prf.cases) |
|
881 |
apply(simp_all) |
|
227 | 882 |
apply(rotate_tac 3) |
883 |
apply(erule Prf.cases) |
|
884 |
apply(simp_all) |
|
885 |
apply(rule Prf.intros) |
|
886 |
apply(auto)[1] |
|
887 |
apply(rotate_tac 3) |
|
888 |
apply(erule Prf.cases) |
|
889 |
apply(simp_all) |
|
890 |
apply(rotate_tac 3) |
|
891 |
apply(erule Prf.cases) |
|
892 |
apply(simp_all) |
|
223 | 893 |
apply(rule Prf.intros) |
894 |
apply(auto)[1] |
|
895 |
apply(rotate_tac 3) |
|
896 |
apply(erule Prf.cases) |
|
897 |
apply(simp_all) |
|
227 | 898 |
apply(rotate_tac 3) |
899 |
apply(erule Prf.cases) |
|
900 |
apply(simp_all) |
|
223 | 901 |
apply(rule Prf.intros) |
227 | 902 |
apply(auto)[2] |
903 |
apply(rotate_tac 3) |
|
904 |
apply(erule Prf.cases) |
|
905 |
apply(simp_all) |
|
906 |
apply(rule Prf.intros) |
|
907 |
apply(auto)[3] |
|
223 | 908 |
apply(rotate_tac 3) |
909 |
apply(erule Prf.cases) |
|
910 |
apply(simp_all) |
|
227 | 911 |
apply(rotate_tac 3) |
912 |
apply(erule Prf.cases) |
|
913 |
apply(simp_all) |
|
914 |
apply(rotate_tac 3) |
|
915 |
apply(erule Prf.cases) |
|
916 |
apply(simp_all) |
|
917 |
apply(rule Prf.intros) |
|
918 |
apply(auto)[3] |
|
919 |
apply(rotate_tac 3) |
|
920 |
apply(erule Prf.cases) |
|
921 |
apply(simp_all) |
|
922 |
apply(erule Prf.cases) |
|
923 |
apply(simp_all) |
|
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
924 |
apply(rotate_tac 3) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
925 |
apply(erule Prf.cases) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
926 |
apply(simp_all) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
927 |
apply(rule Prf.intros) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
928 |
apply(auto) |
243 | 929 |
done |
227 | 930 |
|
222 | 931 |
|
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
932 |
lemma Posix_NTIMES_mkeps: |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
933 |
assumes "[] \<in> r \<rightarrow> mkeps r" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
934 |
shows "[] \<in> NTIMES r n \<rightarrow> Stars (replicate n (mkeps r))" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
935 |
apply(induct n) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
936 |
apply(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
937 |
apply (rule Posix_NTIMES2) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
938 |
apply(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
939 |
apply(subst append_Nil[symmetric]) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
940 |
apply (rule Posix_NTIMES1) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
941 |
apply(auto) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
942 |
apply(rule assms) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
943 |
done |
122
7c6c907660d8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
121
diff
changeset
|
944 |
|
223 | 945 |
lemma Posix_FROMNTIMES_mkeps: |
946 |
assumes "[] \<in> r \<rightarrow> mkeps r" |
|
947 |
shows "[] \<in> FROMNTIMES r n \<rightarrow> Stars (replicate n (mkeps r))" |
|
948 |
apply(induct n) |
|
949 |
apply(simp) |
|
950 |
apply (rule Posix_FROMNTIMES2) |
|
225 | 951 |
apply (rule Posix.intros) |
223 | 952 |
apply(simp) |
953 |
apply(subst append_Nil[symmetric]) |
|
954 |
apply (rule Posix_FROMNTIMES1) |
|
955 |
apply(auto) |
|
956 |
apply(rule assms) |
|
957 |
done |
|
958 |
||
227 | 959 |
lemma Posix_NMTIMES_mkeps: |
960 |
assumes "[] \<in> r \<rightarrow> mkeps r" "n \<le> m" |
|
961 |
shows "[] \<in> NMTIMES r n m \<rightarrow> Stars (replicate n (mkeps r))" |
|
962 |
using assms(2) |
|
963 |
apply(induct n arbitrary: m) |
|
964 |
apply(simp) |
|
965 |
apply(rule Posix.intros) |
|
966 |
apply(rule Posix.intros) |
|
967 |
apply(case_tac m) |
|
968 |
apply(simp) |
|
969 |
apply(simp) |
|
970 |
apply(subst append_Nil[symmetric]) |
|
971 |
apply(rule Posix.intros) |
|
972 |
apply(auto) |
|
973 |
apply(rule assms) |
|
974 |
done |
|
975 |
||
976 |
||
223 | 977 |
|
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
978 |
lemma Posix_mkeps: |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
979 |
assumes "nullable r" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
980 |
shows "[] \<in> r \<rightarrow> mkeps r" |
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
981 |
using assms |
185
841f7b9c0a6a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
172
diff
changeset
|
982 |
apply(induct r rule: nullable.induct) |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
983 |
apply(auto intro: Posix.intros simp add: nullable_correctness Sequ_def) |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
984 |
apply(subst append.simps(1)[symmetric]) |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
985 |
apply(rule Posix.intros) |
123
1bee7006b92b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
122
diff
changeset
|
986 |
apply(auto) |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
987 |
apply(case_tac n) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
988 |
apply(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
989 |
apply (simp add: Posix_NTIMES2) |
223 | 990 |
apply(simp) |
991 |
apply(subst append.simps(1)[symmetric]) |
|
992 |
apply(rule Posix.intros) |
|
993 |
apply(auto) |
|
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
994 |
apply(rule Posix_NTIMES_mkeps) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
995 |
apply(simp) |
225 | 996 |
apply(rule Posix.intros) |
997 |
apply(rule Posix.intros) |
|
223 | 998 |
apply(case_tac n) |
999 |
apply(simp) |
|
225 | 1000 |
apply(rule Posix.intros) |
1001 |
apply(rule Posix.intros) |
|
223 | 1002 |
apply(simp) |
1003 |
apply(subst append.simps(1)[symmetric]) |
|
1004 |
apply(rule Posix.intros) |
|
1005 |
apply(auto) |
|
1006 |
apply(rule Posix_FROMNTIMES_mkeps) |
|
1007 |
apply(simp) |
|
227 | 1008 |
apply(rule Posix.intros) |
1009 |
apply(rule Posix.intros) |
|
1010 |
apply(case_tac n) |
|
1011 |
apply(simp) |
|
1012 |
apply(rule Posix.intros) |
|
1013 |
apply(rule Posix.intros) |
|
1014 |
apply(simp) |
|
1015 |
apply(case_tac m) |
|
1016 |
apply(simp) |
|
1017 |
apply(simp) |
|
1018 |
apply(subst append_Nil[symmetric]) |
|
1019 |
apply(rule Posix.intros) |
|
1020 |
apply(auto) |
|
1021 |
apply(rule Posix_NMTIMES_mkeps) |
|
1022 |
apply(auto) |
|
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1023 |
apply(subst append_Nil[symmetric]) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1024 |
apply(rule Posix.intros) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1025 |
apply(auto) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1026 |
apply(rule Posix.intros) |
91
f067e59b58d9
more lemmas for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
90
diff
changeset
|
1027 |
done |
89
9613e6ace30f
added theory for star
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1028 |
|
142
08dcf0d20f15
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
1029 |
|
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1030 |
lemma Posix_determ: |
122
7c6c907660d8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
121
diff
changeset
|
1031 |
assumes "s \<in> r \<rightarrow> v1" "s \<in> r \<rightarrow> v2" |
7c6c907660d8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
121
diff
changeset
|
1032 |
shows "v1 = v2" |
7c6c907660d8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
121
diff
changeset
|
1033 |
using assms |
151
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1034 |
proof (induct s r v1 arbitrary: v2 rule: Posix.induct) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1035 |
case (Posix_ONE v2) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1036 |
have "[] \<in> ONE \<rightarrow> v2" by fact |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1037 |
then show "Void = v2" by cases auto |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1038 |
next |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1039 |
case (Posix_CHAR c v2) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1040 |
have "[c] \<in> CHAR c \<rightarrow> v2" by fact |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1041 |
then show "Char c = v2" by cases auto |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1042 |
next |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1043 |
case (Posix_ALT1 s r1 v r2 v2) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1044 |
have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1045 |
moreover |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1046 |
have "s \<in> r1 \<rightarrow> v" by fact |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1047 |
then have "s \<in> L r1" by (simp add: Posix1) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1048 |
ultimately obtain v' where eq: "v2 = Left v'" "s \<in> r1 \<rightarrow> v'" by cases auto |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1049 |
moreover |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1050 |
have IH: "\<And>v2. s \<in> r1 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1051 |
ultimately have "v = v'" by simp |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1052 |
then show "Left v = v2" using eq by simp |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1053 |
next |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1054 |
case (Posix_ALT2 s r2 v r1 v2) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1055 |
have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1056 |
moreover |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1057 |
have "s \<notin> L r1" by fact |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1058 |
ultimately obtain v' where eq: "v2 = Right v'" "s \<in> r2 \<rightarrow> v'" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1059 |
by cases (auto simp add: Posix1) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1060 |
moreover |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1061 |
have IH: "\<And>v2. s \<in> r2 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1062 |
ultimately have "v = v'" by simp |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1063 |
then show "Right v = v2" using eq by simp |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1064 |
next |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1065 |
case (Posix_SEQ s1 r1 v1 s2 r2 v2 v') |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1066 |
have "(s1 @ s2) \<in> SEQ r1 r2 \<rightarrow> v'" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1067 |
"s1 \<in> r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1068 |
"\<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+ |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1069 |
then obtain v1' v2' where "v' = Seq v1' v2'" "s1 \<in> r1 \<rightarrow> v1'" "s2 \<in> r2 \<rightarrow> v2'" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1070 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1071 |
using Posix1(1) by fastforce+ |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1072 |
moreover |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1073 |
have IHs: "\<And>v1'. s1 \<in> r1 \<rightarrow> v1' \<Longrightarrow> v1 = v1'" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1074 |
"\<And>v2'. s2 \<in> r2 \<rightarrow> v2' \<Longrightarrow> v2 = v2'" by fact+ |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1075 |
ultimately show "Seq v1 v2 = v'" by simp |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1076 |
next |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1077 |
case (Posix_STAR1 s1 r v s2 vs v2) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1078 |
have "(s1 @ s2) \<in> STAR r \<rightarrow> v2" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1079 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" "flat v \<noteq> []" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1080 |
"\<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+ |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1081 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (STAR r) \<rightarrow> (Stars vs')" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1082 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1083 |
using Posix1(1) apply fastforce |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1084 |
apply (metis Posix1(1) Posix_STAR1.hyps(6) append_Nil append_Nil2) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1085 |
using Posix1(2) by blast |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1086 |
moreover |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1087 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1088 |
"\<And>v2. s2 \<in> STAR r \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1089 |
ultimately show "Stars (v # vs) = v2" by auto |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1090 |
next |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1091 |
case (Posix_STAR2 r v2) |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1092 |
have "[] \<in> STAR r \<rightarrow> v2" by fact |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1093 |
then show "Stars [] = v2" by cases (auto simp add: Posix1) |
220 | 1094 |
next |
1095 |
case (Posix_UPNTIMES1 s1 r v s2 n vs v2) |
|
1096 |
have "(s1 @ s2) \<in> UPNTIMES r (Suc n) \<rightarrow> v2" |
|
1097 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> (UPNTIMES r n) \<rightarrow> Stars vs" "flat v \<noteq> []" |
|
1098 |
"\<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 (UPNTIMES r n))" by fact+ |
|
1099 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (UPNTIMES r n) \<rightarrow> (Stars vs')" |
|
1100 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1101 |
using Posix1(1) apply fastforce |
|
1102 |
apply (metis Posix1(1) Posix_UPNTIMES1.hyps(6) append_Nil append_Nil2) |
|
1103 |
using Posix1(2) by blast |
|
1104 |
moreover |
|
1105 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1106 |
"\<And>v2. s2 \<in> UPNTIMES r n \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1107 |
ultimately show "Stars (v # vs) = v2" by auto |
|
1108 |
next |
|
1109 |
case (Posix_UPNTIMES2 r n v2) |
|
1110 |
have "[] \<in> UPNTIMES r n \<rightarrow> v2" by fact |
|
1111 |
then show "Stars [] = v2" by cases (auto simp add: Posix1) |
|
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1112 |
next |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1113 |
case (Posix_NTIMES2 r v2) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1114 |
have "[] \<in> NTIMES r 0 \<rightarrow> v2" by fact |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1115 |
then show "Stars [] = v2" by cases (auto simp add: Posix1) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1116 |
next |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1117 |
case (Posix_NTIMES1 s1 r v s2 n vs v2) |
236 | 1118 |
have "(s1 @ s2) \<in> NTIMES r (Suc n) \<rightarrow> v2" "flat v = [] \<Longrightarrow> flat (Stars vs) = []" |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1119 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> (NTIMES r n) \<rightarrow> Stars vs" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1120 |
"\<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 (NTIMES r n))" by fact+ |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1121 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (NTIMES r n) \<rightarrow> (Stars vs')" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1122 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1123 |
using Posix1(1) apply fastforce |
236 | 1124 |
by (metis Posix1(1) Posix_NTIMES1.hyps(6) append_Nil append_Nil2) |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1125 |
moreover |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1126 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1127 |
"\<And>v2. s2 \<in> NTIMES r n \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1128 |
ultimately show "Stars (v # vs) = v2" by auto |
223 | 1129 |
next |
225 | 1130 |
case (Posix_FROMNTIMES2 s r v1 v2) |
1131 |
have "s \<in> FROMNTIMES r 0 \<rightarrow> v2" "s \<in> STAR r \<rightarrow> Stars v1" |
|
1132 |
"\<And>v3. s \<in> STAR r \<rightarrow> v3 \<Longrightarrow> Stars v1 = v3" by fact+ |
|
1133 |
then show ?case |
|
1134 |
apply(cases) |
|
1135 |
apply(auto) |
|
1136 |
done |
|
223 | 1137 |
next |
1138 |
case (Posix_FROMNTIMES1 s1 r v s2 n vs v2) |
|
1139 |
have "(s1 @ s2) \<in> FROMNTIMES r (Suc n) \<rightarrow> v2" |
|
236 | 1140 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> (FROMNTIMES r n) \<rightarrow> Stars vs" "flat v = [] \<Longrightarrow> flat (Stars vs) = []" |
223 | 1141 |
"\<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 (FROMNTIMES r n))" by fact+ |
1142 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (FROMNTIMES r n) \<rightarrow> (Stars vs')" |
|
1143 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1144 |
using Posix1(1) apply fastforce |
|
236 | 1145 |
by (metis Posix1(1) Posix_FROMNTIMES1.hyps(6) append_Nil append_Nil2) |
223 | 1146 |
moreover |
1147 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1148 |
"\<And>v2. s2 \<in> FROMNTIMES r n \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1149 |
ultimately show "Stars (v # vs) = v2" by auto |
|
227 | 1150 |
next |
1151 |
case (Posix_NMTIMES2 s r m v1 v2) |
|
1152 |
have "s \<in> NMTIMES r 0 m \<rightarrow> v2" "s \<in> UPNTIMES r m \<rightarrow> Stars v1" |
|
1153 |
"\<And>v3. s \<in> UPNTIMES r m \<rightarrow> v3 \<Longrightarrow> Stars v1 = v3" by fact+ |
|
1154 |
then show ?case |
|
1155 |
apply(cases) |
|
1156 |
apply(auto) |
|
1157 |
done |
|
1158 |
next |
|
1159 |
case (Posix_NMTIMES1 s1 r v s2 n m vs v2) |
|
1160 |
have "(s1 @ s2) \<in> NMTIMES r (Suc n) (Suc m) \<rightarrow> v2" |
|
235 | 1161 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> (NMTIMES r n m) \<rightarrow> Stars vs" "flat v = [] \<Longrightarrow> flat (Stars vs) = []" |
227 | 1162 |
"\<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 (NMTIMES r n m))" by fact+ |
1163 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (NMTIMES r n m) \<rightarrow> (Stars vs')" |
|
1164 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1165 |
using Posix1(1) apply fastforce |
|
235 | 1166 |
by (metis Posix1(1) Posix_NMTIMES1.hyps(6) self_append_conv self_append_conv2) |
227 | 1167 |
moreover |
1168 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1169 |
"\<And>v2. s2 \<in> NMTIMES r n m \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1170 |
ultimately show "Stars (v # vs) = v2" by auto |
|
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1171 |
next |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1172 |
case (Posix_PLUS1 s1 r v s2 vs v2) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1173 |
have "(s1 @ s2) \<in> PLUS r \<rightarrow> v2" |
233
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1174 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> (STAR r) \<rightarrow> Stars vs" (*"flat v = [] \<Longrightarrow> flat (Stars vs) = []"*) |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1175 |
"\<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+ |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1176 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (STAR r) \<rightarrow> (Stars vs')" |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1177 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1178 |
using Posix1(1) apply fastforce |
233
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1179 |
by (metis Posix1(1) Posix_PLUS1.hyps(6) append_self_conv append_self_conv2) |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1180 |
moreover |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1181 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1182 |
"\<And>v2. s2 \<in> STAR r \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1183 |
ultimately show "Stars (v # vs) = v2" by auto |
151
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1184 |
qed |
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1185 |
|
223 | 1186 |
lemma NTIMES_Stars: |
1187 |
assumes "\<turnstile> v : NTIMES r n" |
|
1188 |
shows "\<exists>vs. v = Stars vs \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> length vs = n" |
|
1189 |
using assms |
|
1190 |
apply(induct n arbitrary: v) |
|
1191 |
apply(erule Prf.cases) |
|
1192 |
apply(simp_all) |
|
1193 |
apply(erule Prf.cases) |
|
1194 |
apply(simp_all) |
|
1195 |
done |
|
1196 |
||
1197 |
lemma UPNTIMES_Stars: |
|
1198 |
assumes "\<turnstile> v : UPNTIMES r n" |
|
1199 |
shows "\<exists>vs. v = Stars vs \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> length vs \<le> n" |
|
1200 |
using assms |
|
1201 |
apply(induct n arbitrary: v) |
|
1202 |
apply(erule Prf.cases) |
|
1203 |
apply(simp_all) |
|
1204 |
apply(erule Prf.cases) |
|
1205 |
apply(simp_all) |
|
1206 |
done |
|
1207 |
||
1208 |
lemma FROMNTIMES_Stars: |
|
1209 |
assumes "\<turnstile> v : FROMNTIMES r n" |
|
1210 |
shows "\<exists>vs. v = Stars vs \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> n \<le> length vs" |
|
1211 |
using assms |
|
1212 |
apply(induct n arbitrary: v) |
|
1213 |
apply(erule Prf.cases) |
|
1214 |
apply(simp_all) |
|
1215 |
apply(erule Prf.cases) |
|
1216 |
apply(simp_all) |
|
1217 |
done |
|
1218 |
||
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1219 |
lemma PLUS_Stars: |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1220 |
assumes "\<turnstile> v : PLUS r" |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1221 |
shows "\<exists>vs. v = Stars vs \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> 1 \<le> length vs" |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1222 |
using assms |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1223 |
apply(erule_tac Prf.cases) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1224 |
apply(simp_all) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1225 |
done |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1226 |
|
227 | 1227 |
lemma NMTIMES_Stars: |
1228 |
assumes "\<turnstile> v : NMTIMES r n m" |
|
1229 |
shows "\<exists>vs. v = Stars vs \<and> (\<forall>v \<in> set vs. \<turnstile> v : r) \<and> n \<le> length vs \<and> length vs \<le> m" |
|
1230 |
using assms |
|
1231 |
apply(induct n arbitrary: m v) |
|
1232 |
apply(erule Prf.cases) |
|
1233 |
apply(simp_all) |
|
1234 |
apply(erule Prf.cases) |
|
1235 |
apply(simp_all) |
|
1236 |
done |
|
1237 |
||
100
8b919b3d753e
strengthened PMatch to get determ
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
99
diff
changeset
|
1238 |
|
172
cdc0bdcfba3f
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
151
diff
changeset
|
1239 |
lemma Posix_injval: |
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
1240 |
assumes "s \<in> (der c r) \<rightarrow> v" |
143
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
1241 |
shows "(c # s) \<in> r \<rightarrow> (injval r c v)" |
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
1242 |
using assms |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1243 |
proof(induct r arbitrary: s v rule: rexp.induct) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1244 |
case ZERO |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1245 |
have "s \<in> der c ZERO \<rightarrow> v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1246 |
then have "s \<in> ZERO \<rightarrow> v" by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1247 |
then have "False" by cases |
143
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
1248 |
then show "(c # s) \<in> ZERO \<rightarrow> (injval ZERO c v)" by simp |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1249 |
next |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1250 |
case ONE |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1251 |
have "s \<in> der c ONE \<rightarrow> v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1252 |
then have "s \<in> ZERO \<rightarrow> v" by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1253 |
then have "False" by cases |
143
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
1254 |
then show "(c # s) \<in> ONE \<rightarrow> (injval ONE c v)" by simp |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1255 |
next |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1256 |
case (CHAR d) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1257 |
consider (eq) "c = d" | (ineq) "c \<noteq> d" by blast |
143
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
1258 |
then show "(c # s) \<in> (CHAR d) \<rightarrow> (injval (CHAR d) c v)" |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1259 |
proof (cases) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1260 |
case eq |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1261 |
have "s \<in> der c (CHAR d) \<rightarrow> v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1262 |
then have "s \<in> ONE \<rightarrow> v" using eq by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1263 |
then have eqs: "s = [] \<and> v = Void" by cases simp |
142
08dcf0d20f15
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
1264 |
show "(c # s) \<in> CHAR d \<rightarrow> injval (CHAR d) c v" using eq eqs |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1265 |
by (auto intro: Posix.intros) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1266 |
next |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1267 |
case ineq |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1268 |
have "s \<in> der c (CHAR d) \<rightarrow> v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1269 |
then have "s \<in> ZERO \<rightarrow> v" using ineq by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1270 |
then have "False" by cases |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1271 |
then show "(c # s) \<in> CHAR d \<rightarrow> injval (CHAR d) c v" by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1272 |
qed |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1273 |
next |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1274 |
case (ALT r1 r2) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1275 |
have IH1: "\<And>s v. s \<in> der c r1 \<rightarrow> v \<Longrightarrow> (c # s) \<in> r1 \<rightarrow> injval r1 c v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1276 |
have IH2: "\<And>s v. s \<in> der c r2 \<rightarrow> v \<Longrightarrow> (c # s) \<in> r2 \<rightarrow> injval r2 c v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1277 |
have "s \<in> der c (ALT r1 r2) \<rightarrow> v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1278 |
then have "s \<in> ALT (der c r1) (der c r2) \<rightarrow> v" by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1279 |
then consider (left) v' where "v = Left v'" "s \<in> der c r1 \<rightarrow> v'" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1280 |
| (right) v' where "v = Right v'" "s \<notin> L (der c r1)" "s \<in> der c r2 \<rightarrow> v'" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1281 |
by cases auto |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1282 |
then show "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c v" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1283 |
proof (cases) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1284 |
case left |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1285 |
have "s \<in> der c r1 \<rightarrow> v'" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1286 |
then have "(c # s) \<in> r1 \<rightarrow> injval r1 c v'" using IH1 by simp |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1287 |
then have "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c (Left v')" by (auto intro: Posix.intros) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1288 |
then show "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c v" using left by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1289 |
next |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1290 |
case right |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1291 |
have "s \<notin> L (der c r1)" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1292 |
then have "c # s \<notin> L r1" by (simp add: der_correctness Der_def) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1293 |
moreover |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1294 |
have "s \<in> der c r2 \<rightarrow> v'" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1295 |
then have "(c # s) \<in> r2 \<rightarrow> injval r2 c v'" using IH2 by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1296 |
ultimately have "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c (Right v')" |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1297 |
by (auto intro: Posix.intros) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1298 |
then show "(c # s) \<in> ALT r1 r2 \<rightarrow> injval (ALT r1 r2) c v" using right by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1299 |
qed |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1300 |
next |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1301 |
case (SEQ r1 r2) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1302 |
have IH1: "\<And>s v. s \<in> der c r1 \<rightarrow> v \<Longrightarrow> (c # s) \<in> r1 \<rightarrow> injval r1 c v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1303 |
have IH2: "\<And>s v. s \<in> der c r2 \<rightarrow> v \<Longrightarrow> (c # s) \<in> r2 \<rightarrow> injval r2 c v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1304 |
have "s \<in> der c (SEQ r1 r2) \<rightarrow> v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1305 |
then consider |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1306 |
(left_nullable) v1 v2 s1 s2 where |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1307 |
"v = Left (Seq v1 v2)" "s = s1 @ s2" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1308 |
"s1 \<in> der c r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" "nullable r1" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1309 |
"\<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 (der c r1) \<and> s\<^sub>4 \<in> L r2)" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1310 |
| (right_nullable) v1 s1 s2 where |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1311 |
"v = Right v1" "s = s1 @ s2" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1312 |
"s \<in> der c r2 \<rightarrow> v1" "nullable r1" "s1 @ s2 \<notin> L (SEQ (der c r1) r2)" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1313 |
| (not_nullable) v1 v2 s1 s2 where |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1314 |
"v = Seq v1 v2" "s = s1 @ s2" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1315 |
"s1 \<in> der c r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" "\<not>nullable r1" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1316 |
"\<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 (der c r1) \<and> s\<^sub>4 \<in> L r2)" |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1317 |
by (force split: if_splits elim!: Posix_elims simp add: Sequ_def der_correctness Der_def) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1318 |
then show "(c # s) \<in> SEQ r1 r2 \<rightarrow> injval (SEQ r1 r2) c v" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1319 |
proof (cases) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1320 |
case left_nullable |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1321 |
have "s1 \<in> der c r1 \<rightarrow> v1" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1322 |
then have "(c # s1) \<in> r1 \<rightarrow> injval r1 c v1" using IH1 by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1323 |
moreover |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1324 |
have "\<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 (der c r1) \<and> s\<^sub>4 \<in> L r2)" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1325 |
then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by (simp add: der_correctness Der_def) |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1326 |
ultimately have "((c # s1) @ s2) \<in> SEQ r1 r2 \<rightarrow> Seq (injval r1 c v1) v2" using left_nullable by (rule_tac Posix.intros) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1327 |
then show "(c # s) \<in> SEQ r1 r2 \<rightarrow> injval (SEQ r1 r2) c v" using left_nullable by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1328 |
next |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1329 |
case right_nullable |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1330 |
have "nullable r1" by fact |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1331 |
then have "[] \<in> r1 \<rightarrow> (mkeps r1)" by (rule Posix_mkeps) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1332 |
moreover |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1333 |
have "s \<in> der c r2 \<rightarrow> v1" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1334 |
then have "(c # s) \<in> r2 \<rightarrow> (injval r2 c v1)" using IH2 by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1335 |
moreover |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1336 |
have "s1 @ s2 \<notin> L (SEQ (der c r1) r2)" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1337 |
then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = c # s \<and> [] @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" using right_nullable |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1338 |
by(auto simp add: der_correctness Der_def append_eq_Cons_conv Sequ_def) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1339 |
ultimately have "([] @ (c # s)) \<in> SEQ r1 r2 \<rightarrow> Seq (mkeps r1) (injval r2 c v1)" |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1340 |
by(rule Posix.intros) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1341 |
then show "(c # s) \<in> SEQ r1 r2 \<rightarrow> injval (SEQ r1 r2) c v" using right_nullable by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1342 |
next |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1343 |
case not_nullable |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1344 |
have "s1 \<in> der c r1 \<rightarrow> v1" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1345 |
then have "(c # s1) \<in> r1 \<rightarrow> injval r1 c v1" using IH1 by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1346 |
moreover |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1347 |
have "\<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 (der c r1) \<and> s\<^sub>4 \<in> L r2)" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1348 |
then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by (simp add: der_correctness Der_def) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1349 |
ultimately have "((c # s1) @ s2) \<in> SEQ r1 r2 \<rightarrow> Seq (injval r1 c v1) v2" using not_nullable |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1350 |
by (rule_tac Posix.intros) (simp_all) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1351 |
then show "(c # s) \<in> SEQ r1 r2 \<rightarrow> injval (SEQ r1 r2) c v" using not_nullable by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1352 |
qed |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1353 |
next |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1354 |
case (STAR r) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1355 |
have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1356 |
have "s \<in> der c (STAR r) \<rightarrow> v" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1357 |
then consider |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1358 |
(cons) v1 vs s1 s2 where |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1359 |
"v = Seq v1 (Stars vs)" "s = s1 @ s2" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1360 |
"s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (STAR r) \<rightarrow> (Stars vs)" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1361 |
"\<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 (der c r) \<and> s\<^sub>4 \<in> L (STAR r))" |
149
ec3d221bfc45
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
146
diff
changeset
|
1362 |
apply(auto elim!: Posix_elims(1-5) simp add: der_correctness Der_def intro: Posix.intros) |
142
08dcf0d20f15
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
1363 |
apply(rotate_tac 3) |
149
ec3d221bfc45
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
146
diff
changeset
|
1364 |
apply(erule_tac Posix_elims(6)) |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1365 |
apply (simp add: Posix.intros(6)) |
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1366 |
using Posix.intros(7) by blast |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1367 |
then show "(c # s) \<in> STAR r \<rightarrow> injval (STAR r) c v" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1368 |
proof (cases) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1369 |
case cons |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1370 |
have "s1 \<in> der c r \<rightarrow> v1" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1371 |
then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1372 |
moreover |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1373 |
have "s2 \<in> STAR r \<rightarrow> Stars vs" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1374 |
moreover |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1375 |
have "(c # s1) \<in> r \<rightarrow> injval r c v1" by fact |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1376 |
then have "flat (injval r c v1) = (c # s1)" by (rule Posix1) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1377 |
then have "flat (injval r c v1) \<noteq> []" by simp |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1378 |
moreover |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1379 |
have "\<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 (der c r) \<and> s\<^sub>4 \<in> L (STAR r))" by fact |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1380 |
then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1381 |
by (simp add: der_correctness Der_def) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1382 |
ultimately |
146
da81ffac4b10
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1383 |
have "((c # s1) @ s2) \<in> STAR r \<rightarrow> Stars (injval r c v1 # vs)" by (rule Posix.intros) |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1384 |
then show "(c # s) \<in> STAR r \<rightarrow> injval (STAR r) c v" using cons by(simp) |
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1385 |
qed |
220 | 1386 |
next |
1387 |
case (UPNTIMES r n) |
|
1388 |
have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact |
|
1389 |
have "s \<in> der c (UPNTIMES r n) \<rightarrow> v" by fact |
|
1390 |
then consider |
|
1391 |
(cons) m v1 vs s1 s2 where |
|
1392 |
"n = Suc m" |
|
1393 |
"v = Seq v1 (Stars vs)" "s = s1 @ s2" |
|
1394 |
"s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (UPNTIMES r m) \<rightarrow> (Stars vs)" |
|
1395 |
"\<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 (der c r) \<and> s\<^sub>4 \<in> L (UPNTIMES r m))" |
|
1396 |
apply(case_tac n) |
|
1397 |
apply(simp) |
|
1398 |
using Posix_elims(1) apply blast |
|
1399 |
apply(simp) |
|
223 | 1400 |
apply(auto elim!: Posix_elims(1-5) simp add: der_correctness Der_def intro: Posix.intros) |
220 | 1401 |
by (metis Posix1a UPNTIMES_Stars) |
1402 |
then show "(c # s) \<in> UPNTIMES r n \<rightarrow> injval (UPNTIMES r n) c v" |
|
1403 |
proof (cases) |
|
1404 |
case cons |
|
1405 |
have "n = Suc m" by fact |
|
1406 |
moreover |
|
1407 |
have "s1 \<in> der c r \<rightarrow> v1" by fact |
|
1408 |
then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp |
|
1409 |
moreover |
|
1410 |
have "s2 \<in> UPNTIMES r m \<rightarrow> Stars vs" by fact |
|
1411 |
moreover |
|
1412 |
have "(c # s1) \<in> r \<rightarrow> injval r c v1" by fact |
|
1413 |
then have "flat (injval r c v1) = (c # s1)" by (rule Posix1) |
|
1414 |
then have "flat (injval r c v1) \<noteq> []" by simp |
|
1415 |
moreover |
|
1416 |
have "\<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 (der c r) \<and> s\<^sub>4 \<in> L (UPNTIMES r m))" by fact |
|
1417 |
then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r m))" |
|
1418 |
by (simp add: der_correctness Der_def) |
|
1419 |
ultimately |
|
1420 |
have "((c # s1) @ s2) \<in> UPNTIMES r (Suc m) \<rightarrow> Stars (injval r c v1 # vs)" |
|
1421 |
apply(rule_tac Posix.intros(8)) |
|
1422 |
apply(simp_all) |
|
1423 |
done |
|
1424 |
then show "(c # s) \<in> UPNTIMES r n \<rightarrow> injval (UPNTIMES r n) c v" using cons by(simp) |
|
1425 |
qed |
|
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1426 |
next |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1427 |
case (NTIMES r n) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1428 |
have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1429 |
have "s \<in> der c (NTIMES r n) \<rightarrow> v" by fact |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1430 |
then consider |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1431 |
(cons) m v1 vs s1 s2 where |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1432 |
"n = Suc m" |
236 | 1433 |
"v = Seq v1 (Stars vs)" "s = s1 @ s2" "flat v = [] \<Longrightarrow> flat (Stars vs) = []" |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1434 |
"s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (NTIMES r m) \<rightarrow> (Stars vs)" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1435 |
"\<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 (der c r) \<and> s\<^sub>4 \<in> L (NTIMES r m))" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1436 |
apply(case_tac n) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1437 |
apply(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1438 |
using Posix_elims(1) apply blast |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1439 |
apply(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1440 |
apply(auto elim!: Posix_elims(1-5) simp add: der_correctness Der_def intro: Posix.intros) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1441 |
by (metis NTIMES_Stars Posix1a) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1442 |
then show "(c # s) \<in> NTIMES r n \<rightarrow> injval (NTIMES r n) c v" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1443 |
proof (cases) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1444 |
case cons |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1445 |
have "n = Suc m" by fact |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1446 |
moreover |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1447 |
have "s1 \<in> der c r \<rightarrow> v1" by fact |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1448 |
then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1449 |
moreover |
236 | 1450 |
have "flat v = [] \<Longrightarrow> flat (Stars vs) = []" by fact |
1451 |
moreover |
|
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1452 |
have "s2 \<in> NTIMES r m \<rightarrow> Stars vs" by fact |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1453 |
moreover |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1454 |
have "\<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 (der c r) \<and> s\<^sub>4 \<in> L (NTIMES r m))" by fact |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1455 |
then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (NTIMES r m))" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1456 |
by (simp add: der_correctness Der_def) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1457 |
ultimately |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1458 |
have "((c # s1) @ s2) \<in> NTIMES r (Suc m) \<rightarrow> Stars (injval r c v1 # vs)" |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1459 |
apply(rule_tac Posix.intros(10)) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1460 |
apply(simp_all) |
236 | 1461 |
by (simp add: Posix1(2)) |
221
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1462 |
then show "(c # s) \<in> NTIMES r n \<rightarrow> injval (NTIMES r n) c v" using cons by(simp) |
c21938d0b396
added also the ntimes case
Christian Urban <urbanc@in.tum.de>
parents:
220
diff
changeset
|
1463 |
qed |
223 | 1464 |
next |
1465 |
case (FROMNTIMES r n) |
|
1466 |
have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact |
|
1467 |
have "s \<in> der c (FROMNTIMES r n) \<rightarrow> v" by fact |
|
1468 |
then consider |
|
1469 |
(null) v1 vs s1 s2 where |
|
224 | 1470 |
"n = 0" |
223 | 1471 |
"v = Seq v1 (Stars vs)" "s = s1 @ s2" |
225 | 1472 |
"s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (FROMNTIMES r 0) \<rightarrow> (Stars vs)" |
223 | 1473 |
"\<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 (der c r) \<and> s\<^sub>4 \<in> L (FROMNTIMES r 0))" |
1474 |
| (cons) m v1 vs s1 s2 where |
|
1475 |
"n = Suc m" |
|
236 | 1476 |
"v = Seq v1 (Stars vs)" "s = s1 @ s2" "flat v = [] \<Longrightarrow> flat (Stars vs) = []" |
223 | 1477 |
"s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (FROMNTIMES r m) \<rightarrow> (Stars vs)" |
1478 |
"\<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 (der c r) \<and> s\<^sub>4 \<in> L (FROMNTIMES r m))" |
|
1479 |
apply(case_tac n) |
|
1480 |
apply(simp) |
|
1481 |
apply(auto elim!: Posix_elims(1-5) simp add: der_correctness Der_def intro: Posix.intros) |
|
1482 |
defer |
|
1483 |
apply (metis FROMNTIMES_Stars Posix1a) |
|
1484 |
apply(rotate_tac 5) |
|
225 | 1485 |
apply(erule Posix.cases) |
1486 |
apply(simp_all) |
|
1487 |
apply(clarify) |
|
1488 |
by (simp add: Posix_FROMNTIMES2) |
|
223 | 1489 |
then show "(c # s) \<in> FROMNTIMES r n \<rightarrow> injval (FROMNTIMES r n) c v" |
1490 |
proof (cases) |
|
1491 |
case cons |
|
1492 |
have "n = Suc m" by fact |
|
1493 |
moreover |
|
1494 |
have "s1 \<in> der c r \<rightarrow> v1" by fact |
|
1495 |
then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp |
|
1496 |
moreover |
|
1497 |
have "s2 \<in> FROMNTIMES r m \<rightarrow> Stars vs" by fact |
|
236 | 1498 |
moreover |
1499 |
have "flat v = [] \<Longrightarrow> flat (Stars vs) = []" by fact |
|
223 | 1500 |
moreover |
1501 |
have "\<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 (der c r) \<and> s\<^sub>4 \<in> L (FROMNTIMES r m))" by fact |
|
1502 |
then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (FROMNTIMES r m))" |
|
1503 |
by (simp add: der_correctness Der_def) |
|
1504 |
ultimately |
|
1505 |
have "((c # s1) @ s2) \<in> FROMNTIMES r (Suc m) \<rightarrow> Stars (injval r c v1 # vs)" |
|
1506 |
apply(rule_tac Posix.intros(12)) |
|
236 | 1507 |
apply(simp_all) |
1508 |
by (simp add: Posix1(2)) |
|
223 | 1509 |
then show "(c # s) \<in> FROMNTIMES r n \<rightarrow> injval (FROMNTIMES r n) c v" using cons by(simp) |
224 | 1510 |
next |
1511 |
case null |
|
225 | 1512 |
then have "((c # s1) @ s2) \<in> FROMNTIMES r 0 \<rightarrow> Stars (injval r c v1 # vs)" |
1513 |
apply(rule_tac Posix.intros) |
|
1514 |
apply(rule_tac Posix.intros) |
|
1515 |
apply(rule IH) |
|
224 | 1516 |
apply(simp) |
225 | 1517 |
apply(rotate_tac 4) |
1518 |
apply(erule Posix.cases) |
|
1519 |
apply(simp_all) |
|
1520 |
apply (simp add: Posix1a Prf_injval_flat) |
|
228
8b432cef3805
polished some of the definitions
Christian Urban <urbanc@in.tum.de>
parents:
227
diff
changeset
|
1521 |
apply(simp only: Star_def) |
225 | 1522 |
apply(simp only: der_correctness Der_def) |
1523 |
apply(simp) |
|
1524 |
done |
|
1525 |
then show "(c # s) \<in> FROMNTIMES r n \<rightarrow> injval (FROMNTIMES r n) c v" using null by simp |
|
223 | 1526 |
qed |
227 | 1527 |
next |
1528 |
case (NMTIMES r n m) |
|
1529 |
have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact |
|
1530 |
have "s \<in> der c (NMTIMES r n m) \<rightarrow> v" by fact |
|
1531 |
then show "(c # s) \<in> (NMTIMES r n m) \<rightarrow> injval (NMTIMES r n m) c v" |
|
1532 |
apply(case_tac "m < n") |
|
1533 |
apply(simp) |
|
1534 |
using Posix_elims(1) apply blast |
|
1535 |
apply(case_tac n) |
|
1536 |
apply(simp_all) |
|
1537 |
apply(case_tac m) |
|
1538 |
apply(simp) |
|
1539 |
apply(erule Posix_elims(1)) |
|
1540 |
apply(simp) |
|
1541 |
apply(erule Posix.cases) |
|
1542 |
apply(simp_all) |
|
1543 |
apply(clarify) |
|
1544 |
apply(rotate_tac 5) |
|
1545 |
apply(frule Posix1a) |
|
1546 |
apply(drule UPNTIMES_Stars) |
|
1547 |
apply(clarify) |
|
1548 |
apply(simp) |
|
1549 |
apply(subst append_Cons[symmetric]) |
|
1550 |
apply(rule Posix.intros) |
|
1551 |
apply(rule Posix.intros) |
|
1552 |
apply(auto) |
|
1553 |
apply(rule IH) |
|
1554 |
apply blast |
|
1555 |
using Posix1a Prf_injval_flat apply auto[1] |
|
1556 |
apply(simp add: der_correctness Der_def) |
|
1557 |
apply blast |
|
1558 |
apply(case_tac m) |
|
1559 |
apply(simp) |
|
1560 |
apply(simp) |
|
1561 |
apply(erule Posix.cases) |
|
1562 |
apply(simp_all) |
|
1563 |
apply(clarify) |
|
1564 |
apply(rotate_tac 6) |
|
1565 |
apply(frule Posix1a) |
|
1566 |
apply(drule NMTIMES_Stars) |
|
1567 |
apply(clarify) |
|
1568 |
apply(simp) |
|
1569 |
apply(subst append_Cons[symmetric]) |
|
1570 |
apply(rule Posix.intros) |
|
1571 |
apply(rule IH) |
|
1572 |
apply(blast) |
|
1573 |
apply(simp) |
|
1574 |
apply(simp add: der_correctness Der_def) |
|
235 | 1575 |
using Posix1a Prf_injval_flat list.distinct(1) apply auto[1] |
1576 |
apply(simp add: der_correctness Der_def) |
|
227 | 1577 |
done |
230
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1578 |
next |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1579 |
case (PLUS r) |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1580 |
have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact |
80e7a94f6670
just for fun added the case for PLUS (was already proved as FROMNTIMES)
Christian Urban <urbanc@in.tum.de>
parents:
229
diff
changeset
|
1581 |
have "s \<in> der c (PLUS r) \<rightarrow> v" by fact |
233
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1582 |
then show "(c # s) \<in> PLUS r \<rightarrow> injval (PLUS r) c v" |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1583 |
apply - |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1584 |
apply(erule Posix.cases) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1585 |
apply(simp_all) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1586 |
apply(auto) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1587 |
apply(rotate_tac 3) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1588 |
apply(erule Posix.cases) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1589 |
apply(simp_all) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1590 |
apply(subst append_Cons[symmetric]) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1591 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1592 |
using PLUS.hyps apply auto[1] |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1593 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1594 |
apply(simp) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1595 |
apply(simp) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1596 |
apply(simp) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1597 |
apply(simp) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1598 |
apply(simp) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1599 |
using Posix1a Prf_injval_flat apply auto[1] |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1600 |
apply(simp add: der_correctness Der_def) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1601 |
apply(subst append_Nil2[symmetric]) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1602 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1603 |
using PLUS.hyps apply auto[1] |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1604 |
apply(rule Posix.intros) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1605 |
apply(simp) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1606 |
apply(simp) |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1607 |
done |
121
4c85af262ee7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
120
diff
changeset
|
1608 |
qed |
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
1609 |
|
145
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1610 |
|
227 | 1611 |
|
145
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1612 |
section {* The Lexer by Sulzmann and Lu *} |
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1613 |
|
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1614 |
fun |
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1615 |
lexer :: "rexp \<Rightarrow> string \<Rightarrow> val option" |
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1616 |
where |
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1617 |
"lexer r [] = (if nullable r then Some(mkeps r) else None)" |
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1618 |
| "lexer r (c#s) = (case (lexer (der c r) s) of |
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1619 |
None \<Rightarrow> None |
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1620 |
| Some(v) \<Rightarrow> Some(injval r c v))" |
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1621 |
|
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1622 |
|
151
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1623 |
lemma lexer_correct_None: |
145
97735ef233be
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1624 |
shows "s \<notin> L r \<longleftrightarrow> lexer r s = None" |
120
d74bfa11802c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
113
diff
changeset
|
1625 |
apply(induct s arbitrary: r) |
143
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
1626 |
apply(simp add: nullable_correctness) |
120
d74bfa11802c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
113
diff
changeset
|
1627 |
apply(drule_tac x="der a r" in meta_spec) |
143
1e7b36450d9a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
1628 |
apply(auto simp add: der_correctness Der_def) |
120
d74bfa11802c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
113
diff
changeset
|
1629 |
done |
106
489dfa0d7ec9
more cleaning and moving unnessary stuff to the end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
105
diff
changeset
|
1630 |
|
151
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1631 |
lemma lexer_correct_Some: |
185
841f7b9c0a6a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
172
diff
changeset
|
1632 |
shows "s \<in> L r \<longleftrightarrow> (\<exists>v. lexer r s = Some(v) \<and> s \<in> r \<rightarrow> v)" |
124
5378ddbd1381
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
123
diff
changeset
|
1633 |
apply(induct s arbitrary: r) |
151
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1634 |
apply(auto simp add: Posix_mkeps nullable_correctness)[1] |
124
5378ddbd1381
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
123
diff
changeset
|
1635 |
apply(drule_tac x="der a r" in meta_spec) |
5378ddbd1381
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
123
diff
changeset
|
1636 |
apply(simp add: der_correctness Der_def) |
5378ddbd1381
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
123
diff
changeset
|
1637 |
apply(rule iffI) |
172
cdc0bdcfba3f
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
151
diff
changeset
|
1638 |
apply(auto intro: Posix_injval simp add: Posix1(1)) |
151
5a1196466a9c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
150
diff
changeset
|
1639 |
done |
149
ec3d221bfc45
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
146
diff
changeset
|
1640 |
|
186
0b94800eb616
added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
185
diff
changeset
|
1641 |
lemma lexer_correctness: |
0b94800eb616
added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
185
diff
changeset
|
1642 |
shows "(lexer r s = Some v) \<longleftrightarrow> s \<in> r \<rightarrow> v" |
0b94800eb616
added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
185
diff
changeset
|
1643 |
and "(lexer r s = None) \<longleftrightarrow> \<not>(\<exists>v. s \<in> r \<rightarrow> v)" |
0b94800eb616
added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
185
diff
changeset
|
1644 |
using Posix1(1) Posix_determ lexer_correct_None lexer_correct_Some apply fastforce |
0b94800eb616
added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
185
diff
changeset
|
1645 |
using Posix1(1) lexer_correct_None lexer_correct_Some by blast |
0b94800eb616
added corollary
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
185
diff
changeset
|
1646 |
|
233
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1647 |
value "lexer (PLUS (ALT ONE (SEQ (CHAR (CHR ''a'')) (CHAR (CHR ''b''))))) [CHR ''a'', CHR ''b'']" |
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1648 |
|
654b542ce8db
strengthened PLUS-posix definition
Christian Urban <urbanc@in.tum.de>
parents:
230
diff
changeset
|
1649 |
|
226 | 1650 |
unused_thms |
95
a33d3040bf7e
started a paper and moved cruft to Attic
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
94
diff
changeset
|
1651 |
end |