author | Chengsong |
Mon, 10 Jul 2023 19:29:22 +0100 | |
changeset 664 | ba44144875b1 |
parent 492 | 61eff2abb0b6 |
permissions | -rw-r--r-- |
451 | 1 |
(**) |
444 | 2 |
theory ClosedFormsBounds |
3 |
imports "GeneralRegexBound" "ClosedForms" |
|
4 |
begin |
|
449 | 5 |
lemma alts_ders_lambda_shape_ders: |
6 |
shows "\<forall>r \<in> set (map (\<lambda>r. rders_simp r ( s)) rs ). \<exists>r1 \<in> set rs. r = rders_simp r1 s" |
|
7 |
by (simp add: image_iff) |
|
8 |
||
9 |
lemma rlist_bound: |
|
10 |
shows "\<forall>r \<in> set rs. rsize r \<le> N \<Longrightarrow> sum_list (map rsize rs) \<le> N * (length rs)" |
|
11 |
apply(induct rs) |
|
12 |
apply simp |
|
13 |
by simp |
|
444 | 14 |
|
15 |
||
16 |
lemma alts_closed_form_bounded: shows |
|
17 |
"\<forall>r \<in> set rs. \<forall>s. rsize(rders_simp r s ) \<le> N \<Longrightarrow> |
|
449 | 18 |
rsize (rders_simp (RALTS rs ) s) \<le> max (Suc ( N * (length rs))) (rsize (RALTS rs) )" |
444 | 19 |
apply(induct s) |
20 |
apply simp |
|
449 | 21 |
apply(subst alts_closed_form_variant) |
22 |
apply force |
|
23 |
apply(subgoal_tac "rsize (rsimp (RALTS (map (\<lambda>r. rders_simp r (a # s)) rs))) \<le> rsize ( (RALTS (map (\<lambda>r. rders_simp r (a # s)) rs)))") |
|
447 | 24 |
prefer 2 |
449 | 25 |
using rsimp_mono apply presburger |
26 |
apply(subgoal_tac "rsize (RALTS (map (\<lambda>r. rders_simp r (a # s)) rs)) = |
|
27 |
Suc (sum_list (map rsize (map (\<lambda>r. rders_simp r (a # s)) rs)))") |
|
444 | 28 |
prefer 2 |
449 | 29 |
using rsize.simps(4) apply blast |
30 |
apply(subgoal_tac "sum_list (map rsize (map (\<lambda>r. rders_simp r (a # s)) rs )) \<le> N * (length rs) ") |
|
31 |
apply linarith |
|
32 |
apply(subgoal_tac "\<forall>r \<in> set (map (\<lambda>r. rders_simp r (a # s)) rs ). rsize r \<le> N") |
|
33 |
prefer 2 |
|
34 |
apply(subgoal_tac "\<forall>r \<in> set (map (\<lambda>r. rders_simp r (a # s)) rs ). \<exists>r1 \<in> set rs. r = rders_simp r1 (a # s)") |
|
35 |
prefer 2 |
|
36 |
using alts_ders_lambda_shape_ders apply presburger |
|
37 |
apply metis |
|
38 |
apply(frule rlist_bound) |
|
39 |
by fastforce |
|
444 | 40 |
|
41 |
||
447 | 42 |
lemma alts_simp_ineq_unfold: |
43 |
shows "rsize (rsimp (RALTS rs)) \<le> Suc (sum_list (map rsize (rdistinct (rflts (map rsimp rs)) {})))" |
|
44 |
using rsimp_aalts_smaller by auto |
|
444 | 45 |
|
451 | 46 |
lemma no_elem_distinct : |
47 |
shows "a \<notin> set rs \<Longrightarrow> rdistinct rs rset = rdistinct rs (insert a rset)" |
|
48 |
apply(induct rs arbitrary: rset) |
|
49 |
apply simp |
|
50 |
by force |
|
51 |
||
52 |
||
53 |
||
54 |
||
55 |
inductive good :: "rrexp \<Rightarrow> bool" |
|
56 |
where |
|
57 |
"RZERO \<notin> set rs \<Longrightarrow> good (RALTS rs)" |
|
58 |
|"good RZERO" |
|
59 |
|"good RONE" |
|
60 |
|"good (RCHAR c)" |
|
61 |
|"good (RSEQ r1 r2)" |
|
62 |
|"good (RSTAR r0)" |
|
63 |
||
64 |
||
65 |
lemma after_flts_no0: |
|
66 |
shows "\<forall>r \<in> set rs. good r \<Longrightarrow> RZERO \<notin> set (rflts rs)" |
|
67 |
apply(induct rs) |
|
68 |
apply simp |
|
69 |
apply(case_tac a) |
|
70 |
apply simp |
|
71 |
apply simp |
|
72 |
apply simp |
|
73 |
apply simp |
|
74 |
apply simp |
|
75 |
using good.simps apply blast |
|
76 |
apply simp |
|
77 |
done |
|
78 |
||
447 | 79 |
lemma flts_has_no_zero: |
451 | 80 |
shows "\<forall>r \<in> set rs. good r \<Longrightarrow> rdistinct (rflts rs) rset = rdistinct (rflts rs) (insert RZERO rset)" |
81 |
apply(subgoal_tac "RZERO \<notin> set (rflts rs)") |
|
82 |
apply (meson no_elem_distinct) |
|
83 |
apply(insert after_flts_no0) |
|
84 |
by blast |
|
85 |
||
86 |
lemma shape_of_alts: |
|
87 |
shows "\<forall>a \<in> alts_set. (\<exists>xs. a = RALTS xs) \<Longrightarrow> RZERO \<notin> alts_set \<and> RONE \<notin> alts_set" |
|
88 |
by fastforce |
|
89 |
||
90 |
lemma quantified2_implies1: |
|
91 |
shows "\<forall>r. P \<and> Q \<Longrightarrow> \<forall>r. P" |
|
92 |
by auto |
|
93 |
||
94 |
lemma quantified_quantifiedAB_A: |
|
95 |
shows " (\<forall>a\<in>alts_set. \<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set) \<Longrightarrow> \<forall>a \<in> alts_set. \<exists>xs. a = RALTS xs" |
|
96 |
by fastforce |
|
97 |
||
98 |
||
99 |
||
100 |
||
101 |
lemma list_distinct_removal: |
|
102 |
shows "set rs \<subseteq> rset \<Longrightarrow> rdistinct (rs @ rs1) rset = rdistinct rs1 rset" |
|
103 |
apply(induct rs arbitrary: rset) |
|
104 |
apply simp |
|
105 |
by simp |
|
106 |
||
107 |
||
108 |
lemma rdistinct_mono_list: |
|
109 |
shows " sum_list (map rsize (rdistinct (x5 @ rs) rset )) \<le> |
|
110 |
( sum_list (map rsize x5)) + (sum_list (map rsize (rdistinct rs ((set x5 ) \<union> rset))))" |
|
111 |
apply(induct x5 arbitrary: rs rset) |
|
112 |
apply simp |
|
113 |
apply(case_tac "a \<in> rset") |
|
114 |
apply simp |
|
115 |
apply (simp add: add.assoc insert_absorb trans_le_add2) |
|
116 |
apply simp |
|
117 |
by (metis Un_insert_right) |
|
118 |
||
119 |
||
120 |
lemma flts_size_reduction_alts: |
|
121 |
shows " \<And>a rs noalts_set alts_set corr_set x5. |
|
122 |
\<lbrakk>\<And>noalts_set alts_set corr_set. |
|
123 |
(\<forall>r\<in>noalts_set. \<forall>xs. r \<noteq> RALTS xs) \<and> |
|
124 |
(\<forall>a\<in>alts_set. \<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set) \<Longrightarrow> |
|
125 |
Suc (sum_list (map rsize (rdistinct (rflts rs) (noalts_set \<union> corr_set)))) |
|
126 |
\<le> Suc (sum_list (map rsize (rdistinct rs (insert RZERO (noalts_set \<union> alts_set))))); |
|
127 |
(\<forall>r\<in>noalts_set. \<forall>xs. r \<noteq> RALTS xs) \<and> |
|
128 |
(\<forall>a\<in>alts_set. \<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set); |
|
129 |
a = RALTS x5\<rbrakk> |
|
130 |
\<Longrightarrow> Suc (sum_list (map rsize (rdistinct (rflts (a # rs)) (noalts_set \<union> corr_set)))) |
|
131 |
\<le> Suc (sum_list |
|
132 |
(map rsize (rdistinct (a # rs) (insert RZERO (noalts_set \<union> alts_set)))))" |
|
133 |
apply(case_tac "a \<in> alts_set") |
|
134 |
apply simp |
|
135 |
apply(subgoal_tac "set x5 \<subseteq> corr_set") |
|
136 |
apply(subst list_distinct_removal) |
|
137 |
apply auto[1] |
|
138 |
apply presburger |
|
139 |
apply fastforce |
|
140 |
apply (subgoal_tac "a \<notin> noalts_set") |
|
141 |
prefer 2 |
|
142 |
apply blast |
|
143 |
apply simp |
|
144 |
apply(subgoal_tac "sum_list (map rsize (rdistinct (x5 @ rflts rs) (noalts_set \<union> corr_set))) |
|
145 |
\<le> sum_list (map rsize x5 ) + sum_list (map rsize (rdistinct (rflts rs) ((set x5) \<union> (noalts_set \<union> corr_set))))") |
|
146 |
prefer 2 |
|
147 |
using rdistinct_mono_list apply presburger |
|
148 |
apply(subgoal_tac "insert (RALTS x5) (noalts_set \<union> alts_set) = noalts_set \<union> (insert (RALTS x5) alts_set)") |
|
149 |
apply(simp only:) |
|
150 |
apply(subgoal_tac "sum_list (map rsize x5) + |
|
151 |
sum_list (map rsize (rdistinct (rflts rs) (noalts_set \<union> (corr_set \<union> (set x5))))) \<le> |
|
152 |
Suc (sum_list (map rsize x5) + |
|
153 |
sum_list |
|
154 |
(map rsize |
|
155 |
(rdistinct rs (insert RZERO (noalts_set \<union> insert (RALTS x5) alts_set)))))") |
|
156 |
||
157 |
apply (simp add: Un_left_commute inf_sup_aci(5)) |
|
158 |
apply(subgoal_tac "sum_list (map rsize (rdistinct (rflts rs) (noalts_set \<union> (corr_set \<union> set x5)))) |
|
159 |
\<le> sum_list |
|
160 |
(map rsize |
|
161 |
(rdistinct rs (insert RZERO (noalts_set \<union> insert (RALTS x5) alts_set))))") |
|
162 |
apply linarith |
|
163 |
apply(subgoal_tac "\<forall>r \<in> insert (RALTS x5) alts_set. \<exists>xs1.( r = RALTS xs1 \<and> set xs1 \<subseteq> corr_set \<union> set x5)") |
|
164 |
apply presburger |
|
165 |
apply (meson insert_iff sup.cobounded2 sup.coboundedI1) |
|
166 |
by blast |
|
167 |
||
168 |
||
169 |
||
449 | 170 |
|
451 | 171 |
lemma flts_vs_nflts1: |
172 |
shows "(\<forall>r \<in> noalts_set. \<forall>xs. r \<noteq> RALTS xs) |
|
173 |
\<and> (\<forall>a \<in> alts_set. (\<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set)) |
|
174 |
\<Longrightarrow> (sum_list (map rsize (rdistinct ( rflts rs) (noalts_set \<union> corr_set) ))) |
|
175 |
\<le> (sum_list (map rsize (rdistinct rs (insert RZERO (noalts_set \<union> alts_set) ))))" |
|
176 |
apply(induct rs arbitrary: noalts_set alts_set corr_set) |
|
177 |
apply simp |
|
178 |
apply(case_tac a) |
|
179 |
apply(case_tac "RZERO \<in> noalts_set") |
|
180 |
apply simp |
|
181 |
apply(subgoal_tac "RZERO \<notin> alts_set") |
|
182 |
apply simp |
|
183 |
apply fastforce |
|
184 |
apply(case_tac "RONE \<in> noalts_set") |
|
185 |
apply simp |
|
186 |
apply(subgoal_tac "RONE \<notin> alts_set") |
|
187 |
prefer 2 |
|
188 |
apply fastforce |
|
189 |
apply(case_tac "RONE \<in> corr_set") |
|
190 |
apply(subgoal_tac "rflts (a # rs) = RONE # rflts rs") |
|
191 |
apply(simp only:) |
|
192 |
apply(subgoal_tac "rdistinct (RONE # rflts rs) (noalts_set \<union> corr_set) = |
|
193 |
rdistinct (rflts rs) (noalts_set \<union> corr_set)") |
|
194 |
apply(simp only:) |
|
195 |
apply(subgoal_tac "rdistinct (RONE # rs) (insert RZERO (noalts_set \<union> alts_set)) = |
|
196 |
RONE # (rdistinct rs (insert RONE (insert RZERO (noalts_set \<union> alts_set)))) ") |
|
197 |
apply(simp only:) |
|
198 |
apply(subgoal_tac "rdistinct (rflts rs) (noalts_set \<union> corr_set) = |
|
199 |
rdistinct (rflts rs) (insert RONE (noalts_set \<union> corr_set))") |
|
200 |
apply (simp only:) |
|
201 |
apply(subgoal_tac "insert RONE (noalts_set \<union> corr_set) = (insert RONE noalts_set) \<union> corr_set") |
|
202 |
apply(simp only:) |
|
203 |
apply(subgoal_tac "insert RONE (insert RZERO (noalts_set \<union> alts_set)) = |
|
204 |
insert RZERO ((insert RONE noalts_set) \<union> alts_set)") |
|
205 |
apply(simp only:) |
|
206 |
apply(subgoal_tac " (sum_list |
|
207 |
(map rsize ( rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set))))) |
|
208 |
\<le> (sum_list |
|
209 |
(map rsize (RONE # rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set)))))") |
|
210 |
apply (smt (verit, best) dual_order.trans insert_iff rrexp.distinct(15)) |
|
480 | 211 |
apply (metis (no_types, opaque_lifting) le_add_same_cancel2 list.simps(9) sum_list.Cons zero_le) |
451 | 212 |
apply fastforce |
213 |
apply fastforce |
|
214 |
apply (metis Un_iff insert_absorb) |
|
215 |
apply (metis UnE insertE insert_is_Un rdistinct.simps(2) rrexp.distinct(1)) |
|
216 |
apply (meson UnCI rdistinct.simps(2)) |
|
217 |
using rflts.simps(4) apply presburger |
|
218 |
apply simp |
|
219 |
apply(subgoal_tac "insert RONE (noalts_set \<union> corr_set) = (insert RONE noalts_set) \<union> corr_set") |
|
220 |
apply(simp only:) |
|
221 |
apply (metis Un_insert_left insertE rrexp.distinct(15)) |
|
222 |
apply fastforce |
|
223 |
apply(case_tac "a \<in> noalts_set") |
|
224 |
apply simp |
|
225 |
apply(subgoal_tac "a \<notin> alts_set") |
|
226 |
prefer 2 |
|
227 |
apply blast |
|
228 |
apply(case_tac "a \<in> corr_set") |
|
229 |
apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set \<union> corr_set)") |
|
230 |
prefer 2 |
|
231 |
apply fastforce |
|
232 |
apply(simp only:) |
|
233 |
apply(subgoal_tac " sum_list |
|
234 |
(map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set)))) |
|
235 |
\<le> |
|
236 |
sum_list |
|
237 |
(map rsize (rdistinct (a # rs) (insert RZERO (noalts_set \<union> alts_set))))") |
|
238 |
||
239 |
apply(subgoal_tac |
|
240 |
"sum_list (map rsize (rdistinct (rflts (a # rs)) ((insert a noalts_set) \<union> corr_set))) |
|
241 |
\<le> |
|
242 |
sum_list (map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))))") |
|
243 |
apply fastforce |
|
244 |
apply simp |
|
245 |
apply(subgoal_tac "(insert a (noalts_set \<union> alts_set)) = (insert a noalts_set) \<union> alts_set") |
|
246 |
apply(simp only:) |
|
247 |
apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set") |
|
248 |
apply(simp only:) |
|
249 |
apply (metis insertE rrexp.distinct(21)) |
|
250 |
apply blast |
|
251 |
||
252 |
apply fastforce |
|
253 |
apply force |
|
254 |
apply simp |
|
255 |
apply (metis Un_insert_left insert_iff rrexp.distinct(21)) |
|
256 |
apply(case_tac "a \<in> noalts_set") |
|
257 |
apply simp |
|
258 |
apply(subgoal_tac "a \<notin> alts_set") |
|
259 |
prefer 2 |
|
260 |
apply blast |
|
261 |
apply(case_tac "a \<in> corr_set") |
|
262 |
apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set \<union> corr_set)") |
|
263 |
prefer 2 |
|
264 |
apply fastforce |
|
265 |
apply(simp only:) |
|
266 |
apply(subgoal_tac " sum_list |
|
267 |
(map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set)))) |
|
268 |
\<le> |
|
269 |
sum_list |
|
270 |
(map rsize (rdistinct (a # rs) (insert RZERO (noalts_set \<union> alts_set))))") |
|
271 |
||
272 |
apply(subgoal_tac |
|
273 |
"sum_list (map rsize (rdistinct (rflts (a # rs)) ((insert a noalts_set) \<union> corr_set))) |
|
274 |
\<le> |
|
275 |
sum_list (map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))))") |
|
276 |
apply fastforce |
|
277 |
apply simp |
|
278 |
apply(subgoal_tac "(insert a (noalts_set \<union> alts_set)) = (insert a noalts_set) \<union> alts_set") |
|
279 |
apply(simp only:) |
|
280 |
apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set") |
|
281 |
apply(simp only:) |
|
282 |
||
283 |
||
284 |
apply (metis insertE rrexp.distinct(25)) |
|
285 |
apply blast |
|
286 |
apply fastforce |
|
287 |
apply force |
|
288 |
apply simp |
|
289 |
||
290 |
apply (metis Un_insert_left insertE rrexp.distinct(25)) |
|
291 |
||
292 |
using Suc_le_mono flts_size_reduction_alts apply presburger |
|
293 |
apply(case_tac "a \<in> noalts_set") |
|
294 |
apply simp |
|
295 |
apply(subgoal_tac "a \<notin> alts_set") |
|
296 |
prefer 2 |
|
297 |
apply blast |
|
298 |
apply(case_tac "a \<in> corr_set") |
|
299 |
apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set \<union> corr_set)") |
|
300 |
prefer 2 |
|
301 |
apply fastforce |
|
302 |
apply(simp only:) |
|
303 |
apply(subgoal_tac " sum_list |
|
304 |
(map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set)))) |
|
305 |
\<le> |
|
306 |
sum_list |
|
307 |
(map rsize (rdistinct (a # rs) (insert RZERO (noalts_set \<union> alts_set))))") |
|
308 |
||
309 |
apply(subgoal_tac |
|
310 |
"sum_list (map rsize (rdistinct (rflts (a # rs)) ((insert a noalts_set) \<union> corr_set))) |
|
311 |
\<le> |
|
312 |
sum_list (map rsize (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))))") |
|
313 |
apply fastforce |
|
314 |
apply simp |
|
315 |
apply(subgoal_tac "(insert a (noalts_set \<union> alts_set)) = (insert a noalts_set) \<union> alts_set") |
|
316 |
apply(simp only:) |
|
317 |
apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set") |
|
318 |
apply(simp only:) |
|
319 |
apply (metis insertE rrexp.distinct(29)) |
|
320 |
||
321 |
apply blast |
|
322 |
||
323 |
apply fastforce |
|
324 |
apply force |
|
325 |
apply simp |
|
326 |
apply (metis Un_insert_left insert_iff rrexp.distinct(29)) |
|
327 |
done |
|
328 |
||
329 |
||
330 |
||
331 |
||
447 | 332 |
|
333 |
lemma flts_vs_nflts: |
|
451 | 334 |
shows "(\<forall>r \<in> noalts_set. \<forall>xs. r \<noteq> RALTS xs) |
335 |
\<and> (\<forall>a \<in> alts_set. (\<exists>xs. a = RALTS xs \<and> set xs \<subseteq> corr_set)) |
|
447 | 336 |
\<Longrightarrow> Suc (sum_list (map rsize (rdistinct ( rflts rs) (noalts_set \<union> corr_set) ))) |
451 | 337 |
\<le> Suc (sum_list (map rsize (rdistinct rs (insert RZERO (noalts_set \<union> alts_set) ))))" |
338 |
apply(induct rs arbitrary: noalts_set alts_set corr_set) |
|
447 | 339 |
apply simp |
451 | 340 |
apply(case_tac a) |
341 |
apply(case_tac "RZERO \<in> noalts_set") |
|
342 |
apply simp |
|
343 |
apply(subgoal_tac "RZERO \<notin> alts_set") |
|
344 |
apply simp |
|
345 |
apply fastforce |
|
346 |
apply(case_tac "RONE \<in> noalts_set") |
|
347 |
apply simp |
|
348 |
apply(subgoal_tac "RONE \<notin> alts_set") |
|
349 |
prefer 2 |
|
350 |
apply fastforce |
|
351 |
apply(case_tac "RONE \<in> corr_set") |
|
352 |
apply(subgoal_tac "rflts (a # rs) = RONE # rflts rs") |
|
353 |
apply(simp only:) |
|
354 |
apply(subgoal_tac "rdistinct (RONE # rflts rs) (noalts_set \<union> corr_set) = |
|
355 |
rdistinct (rflts rs) (noalts_set \<union> corr_set)") |
|
356 |
apply(simp only:) |
|
357 |
apply(subgoal_tac "rdistinct (RONE # rs) (insert RZERO (noalts_set \<union> alts_set)) = |
|
358 |
RONE # (rdistinct rs (insert RONE (insert RZERO (noalts_set \<union> alts_set)))) ") |
|
359 |
apply(simp only:) |
|
360 |
apply(subgoal_tac "rdistinct (rflts rs) (noalts_set \<union> corr_set) = |
|
361 |
rdistinct (rflts rs) (insert RONE (noalts_set \<union> corr_set))") |
|
362 |
apply (simp only:) |
|
363 |
apply(subgoal_tac "insert RONE (noalts_set \<union> corr_set) = (insert RONE noalts_set) \<union> corr_set") |
|
364 |
apply(simp only:) |
|
365 |
apply(subgoal_tac "insert RONE (insert RZERO (noalts_set \<union> alts_set)) = |
|
366 |
insert RZERO ((insert RONE noalts_set) \<union> alts_set)") |
|
367 |
apply(simp only:) |
|
368 |
apply(subgoal_tac " Suc (sum_list |
|
369 |
(map rsize ( rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set))))) |
|
370 |
\<le> Suc (sum_list |
|
371 |
(map rsize (RONE # rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set)))))") |
|
372 |
apply (smt (verit, best) dual_order.trans insert_iff rrexp.distinct(15)) |
|
373 |
apply (metis (no_types, opaque_lifting) add_Suc_right le_add_same_cancel2 list.simps(9) sum_list.Cons zero_le) |
|
374 |
apply fastforce |
|
375 |
apply fastforce |
|
376 |
apply (metis Un_iff insert_absorb) |
|
377 |
apply (metis UnE insertE insert_is_Un rdistinct.simps(2) rrexp.distinct(1)) |
|
378 |
apply (meson UnCI rdistinct.simps(2)) |
|
379 |
using rflts.simps(4) apply presburger |
|
380 |
apply simp |
|
381 |
apply(subgoal_tac "insert RONE (noalts_set \<union> corr_set) = (insert RONE noalts_set) \<union> corr_set") |
|
382 |
apply(simp only:) |
|
383 |
apply (metis Un_insert_left insertE rrexp.distinct(15)) |
|
384 |
apply fastforce |
|
385 |
apply(case_tac "a \<in> noalts_set") |
|
386 |
apply simp |
|
387 |
apply(subgoal_tac "a \<notin> alts_set") |
|
388 |
prefer 2 |
|
389 |
apply blast |
|
390 |
apply(case_tac "a \<in> corr_set") |
|
391 |
apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set \<union> corr_set)") |
|
392 |
prefer 2 |
|
393 |
apply fastforce |
|
394 |
apply(simp only:) |
|
395 |
apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc) |
|
396 |
apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc) |
|
397 |
apply(case_tac "a \<in> noalts_set") |
|
398 |
apply simp |
|
399 |
apply(subgoal_tac "a \<notin> alts_set") |
|
400 |
prefer 2 |
|
401 |
apply blast |
|
402 |
apply(case_tac "a \<in> corr_set") |
|
403 |
apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set \<union> corr_set)") |
|
404 |
prefer 2 |
|
405 |
apply fastforce |
|
406 |
apply(simp only:) |
|
407 |
apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc) |
|
408 |
apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc) |
|
409 |
prefer 2 |
|
410 |
apply(case_tac "a \<in> noalts_set") |
|
411 |
apply simp |
|
412 |
apply(subgoal_tac "a \<notin> alts_set") |
|
413 |
prefer 2 |
|
414 |
apply blast |
|
415 |
apply(case_tac "a \<in> corr_set") |
|
416 |
apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set \<union> corr_set)") |
|
417 |
prefer 2 |
|
418 |
apply fastforce |
|
419 |
apply(simp only:) |
|
420 |
apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc) |
|
421 |
apply (metis add_mono_thms_linordered_semiring(1) flts_vs_nflts1 le_numeral_extra(4) plus_1_eq_Suc) |
|
422 |
using flts_size_reduction_alts apply presburger |
|
423 |
done |
|
447 | 424 |
|
451 | 425 |
|
426 |
||
427 |
||
428 |
||
429 |
(* apply(rutac x = "\<lambda>a rs noalts_set alts_set corr_set. set xs ")*) |
|
430 |
||
447 | 431 |
|
449 | 432 |
lemma distinct_simp_ineq_general: |
450 | 433 |
shows "rsimp ` no_simp = has_simp \<and> finite no_simp \<Longrightarrow>Suc (sum_list (map rsize (rdistinct (map rsimp rs) has_simp))) |
449 | 434 |
\<le> Suc (sum_list (map rsize (rdistinct rs no_simp)))" |
450 | 435 |
apply(induct rs arbitrary: no_simp has_simp) |
436 |
apply simp |
|
437 |
apply(case_tac "a \<in> no_simp") |
|
438 |
apply(subgoal_tac "rsimp a \<in> has_simp") |
|
439 |
apply auto[1] |
|
440 |
apply blast |
|
441 |
apply(case_tac "rsimp a \<in> no_simp") |
|
442 |
apply(subgoal_tac "rsimp a \<in> has_simp") |
|
443 |
prefer 2 |
|
451 | 444 |
apply (simp add: rev_image_eqI rsimp_idem) |
445 |
apply simp |
|
446 |
apply (metis finite_insert image_insert insert_absorb trans_le_add2) |
|
447 |
apply(case_tac "rsimp a \<notin> has_simp") |
|
448 |
apply simp |
|
449 |
apply (metis add_mono_thms_linordered_semiring(1) finite.insertI image_insert rsimp_mono) |
|
450 |
apply simp |
|
451 |
by (metis finite.insertI image_insert insert_absorb trans_le_add2) |
|
449 | 452 |
|
447 | 453 |
|
450 | 454 |
lemma not_mentioned_elem_distinct_strong: |
455 |
shows "r \<noteq> a \<Longrightarrow> (r \<in> set (rdistinct rs rset)) = (r \<in> set (rdistinct rs (insert a rset)))" |
|
456 |
apply(induct rs arbitrary: rset) |
|
457 |
apply simp |
|
458 |
by force |
|
459 |
||
460 |
lemma not_mentioned_elem_distinct: |
|
461 |
shows "r \<noteq> a \<Longrightarrow> (r \<in> set (rdistinct rs {})) = (r \<in> set (rdistinct rs (insert a {})))" |
|
462 |
by (metis not_mentioned_elem_distinct_strong) |
|
463 |
||
464 |
||
465 |
||
466 |
||
451 | 467 |
lemma larger_acc_smaller_distinct_res0: |
468 |
shows " ss \<subseteq> SS \<Longrightarrow> sum_list (map rsize (rdistinct rs SS)) \<le> sum_list (map rsize (rdistinct rs ss))" |
|
469 |
apply(induct rs arbitrary: ss SS) |
|
470 |
apply simp |
|
471 |
apply(case_tac "a \<in> ss") |
|
472 |
apply(subgoal_tac "a \<in> SS") |
|
473 |
apply simp |
|
474 |
apply blast |
|
475 |
apply(case_tac "a \<in> SS") |
|
476 |
apply simp |
|
477 |
apply(subgoal_tac "insert a ss \<subseteq> SS") |
|
478 |
apply simp |
|
479 |
apply (simp add: trans_le_add2) |
|
480 |
apply blast |
|
481 |
apply(simp) |
|
482 |
apply(subgoal_tac "insert a ss \<subseteq> insert a SS") |
|
483 |
apply blast |
|
484 |
by blast |
|
450 | 485 |
|
486 |
||
447 | 487 |
lemma without_flts_ineq: |
488 |
shows " Suc (sum_list (map rsize (rdistinct (rflts rs) {}) )) \<le> |
|
489 |
Suc (sum_list (map rsize (rdistinct ( rs ) {} )))" |
|
451 | 490 |
proof - |
491 |
have " Suc (sum_list (map rsize (rdistinct (rflts rs) {}) )) \<le> |
|
492 |
Suc (sum_list (map rsize (rdistinct rs (insert RZERO {}))))" |
|
493 |
by (metis empty_iff flts_vs_nflts sup_bot_left) |
|
494 |
also have "... \<le> Suc (sum_list (map rsize (rdistinct rs {})))" |
|
495 |
by (simp add: larger_acc_smaller_distinct_res0) |
|
452
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
496 |
finally show ?thesis |
451 | 497 |
by blast |
498 |
qed |
|
499 |
||
500 |
||
447 | 501 |
|
502 |
||
448 | 503 |
|
504 |
||
447 | 505 |
lemma distinct_simp_ineq: |
506 |
shows "Suc (sum_list (map rsize (rdistinct (map rsimp rs) {}))) |
|
507 |
\<le> Suc (sum_list (map rsize (rdistinct rs {})))" |
|
448 | 508 |
|
509 |
using distinct_simp_ineq_general by blast |
|
510 |
||
511 |
||
447 | 512 |
|
513 |
||
514 |
lemma alts_simp_control: |
|
515 |
shows "rsize (rsimp (RALTS rs)) \<le> Suc (sum_list (map rsize (rdistinct rs {})))" |
|
516 |
proof - |
|
517 |
have "rsize (rsimp (RALTS rs)) \<le> Suc (sum_list (map rsize (rdistinct (rflts (map rsimp rs)) {})))" |
|
452
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
518 |
using alts_simp_ineq_unfold by auto |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
519 |
moreover have "\<dots> \<le> Suc (sum_list (map rsize (rdistinct (map rsimp rs) {})))" |
447 | 520 |
using without_flts_ineq by blast |
452
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
521 |
ultimately show "rsize (rsimp (RALTS rs)) \<le> Suc (sum_list (map rsize (rdistinct rs {})))" |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
522 |
by (meson distinct_simp_ineq order.trans) |
447 | 523 |
qed |
444 | 524 |
|
525 |
||
526 |
||
446 | 527 |
lemma rdistinct_equality1: |
528 |
shows "a \<notin> ss \<Longrightarrow> rdistinct (a # rs) ss = a # rdistinct rs (insert a ss) " |
|
529 |
by auto |
|
530 |
||
451 | 531 |
|
446 | 532 |
|
533 |
||
534 |
lemma larger_acc_smaller_distinct_res: |
|
535 |
shows " (sum_list (map rsize (rdistinct rs ss))) \<ge> (sum_list (map rsize (rdistinct rs (insert a ss))))" |
|
447 | 536 |
apply(subgoal_tac "ss \<subseteq> (insert a ss)") |
537 |
apply(rule larger_acc_smaller_distinct_res0) |
|
538 |
apply simp |
|
539 |
by (simp add: subset_insertI) |
|
446 | 540 |
|
541 |
lemma size_list_triangle1: |
|
542 |
shows "sum_list (map rsize (a # (rdistinct as ss))) \<ge> rsize a + sum_list (map rsize (rdistinct as (insert a ss)))" |
|
543 |
by (simp add: larger_acc_smaller_distinct_res) |
|
544 |
||
545 |
||
445 | 546 |
lemma triangle_inequality_distinct: |
547 |
shows "sum_list (map rsize (rdistinct (a # rs) ss)) \<le> rsize a + (sum_list (map rsize (rdistinct rs ss)))" |
|
548 |
apply(case_tac "a \<in> ss") |
|
549 |
apply simp |
|
446 | 550 |
apply(subst rdistinct_equality1) |
551 |
apply simp |
|
552 |
using size_list_triangle1 by auto |
|
445 | 553 |
|
554 |
lemma same_regex_property_after_map: |
|
555 |
shows "\<forall>s. P (f r2 s) \<Longrightarrow> \<forall>r \<in> set (map (f r2) Ss). P r" |
|
556 |
by auto |
|
557 |
||
558 |
lemma same_property_after_distinct: |
|
559 |
shows " \<forall>r \<in> set (map (f r2) Ss). P r \<Longrightarrow> \<forall>r \<in> set (rdistinct (map (f r2) Ss) xset). P r" |
|
560 |
apply(induct Ss arbitrary: xset) |
|
561 |
apply simp |
|
562 |
by auto |
|
563 |
||
564 |
lemma same_regex_property_after_distinct: |
|
565 |
shows "\<forall>s. P (f r2 s) \<Longrightarrow> \<forall>r \<in> set (rdistinct (map (f r2) Ss) xset). P r" |
|
566 |
apply(rule same_property_after_distinct) |
|
567 |
apply(rule same_regex_property_after_map) |
|
568 |
by simp |
|
569 |
||
449 | 570 |
|
571 |
||
572 |
lemma Sum_cons: |
|
573 |
shows "distinct (a # as) \<Longrightarrow> \<Sum> (set ((a::nat) # as)) = a + \<Sum> (set as)" |
|
574 |
by simp |
|
575 |
||
576 |
||
577 |
lemma distinct_list_sizeNregex_bounded: |
|
578 |
shows "distinct rs \<and> (\<forall> r \<in> (set rs). rsize r \<le> N) \<Longrightarrow> sum_list (map rsize rs) \<le> N * length rs" |
|
579 |
apply(induct rs) |
|
580 |
apply simp |
|
581 |
by simp |
|
582 |
||
583 |
||
584 |
lemma distinct_list_size_len_bounded: |
|
585 |
shows "distinct rs \<and> (\<forall>r \<in> set rs. rsize r \<le> N) \<and> length rs \<le> lrs \<Longrightarrow> sum_list (map rsize rs) \<le> lrs * N " |
|
586 |
by (metis distinct_list_sizeNregex_bounded dual_order.trans mult.commute mult_le_mono1) |
|
587 |
||
588 |
||
589 |
||
590 |
lemma rdistinct_same_set: |
|
591 |
shows "(r \<in> set rs) = (r \<in> set (rdistinct rs {}))" |
|
592 |
apply(induct rs) |
|
593 |
apply simp |
|
594 |
apply(case_tac "a \<in> set rs") |
|
595 |
apply(case_tac "r = a") |
|
596 |
apply (simp) |
|
597 |
apply (simp add: not_mentioned_elem_distinct) |
|
598 |
using not_mentioned_elem_distinct by fastforce |
|
599 |
||
600 |
||
601 |
||
602 |
lemma distinct_list_rexp_up_to_certain_size_bouded_by_set_enumerating_up_to_that_size: |
|
603 |
shows "\<forall>r\<in> set rs. (rsize r ) \<le> N \<Longrightarrow> sum_list (map rsize (rdistinct rs {})) \<le> |
|
604 |
(card (sizeNregex N))* N" |
|
605 |
apply(subgoal_tac "distinct (rdistinct rs {})") |
|
606 |
prefer 2 |
|
607 |
using rdistinct_does_the_job apply blast |
|
608 |
apply(subgoal_tac "length (rdistinct rs {}) \<le> card (sizeNregex N)") |
|
609 |
apply(rule distinct_list_size_len_bounded) |
|
610 |
apply(rule conjI)+ |
|
611 |
apply simp |
|
612 |
apply(rule conjI) |
|
613 |
apply (meson rdistinct_same_set) |
|
614 |
apply blast |
|
615 |
apply(subgoal_tac "\<forall>r \<in> set (rdistinct rs {}). rsize r \<le> N") |
|
616 |
prefer 2 |
|
617 |
apply (meson rdistinct_same_set) |
|
618 |
apply(subgoal_tac "length (rdistinct rs {}) = card (set (rdistinct rs {}))") |
|
619 |
prefer 2 |
|
620 |
using set_related_list apply blast |
|
621 |
apply(simp only:) |
|
622 |
by (metis card_mono finite_size_n mem_Collect_eq sizeNregex_def subset_code(1)) |
|
623 |
||
624 |
||
625 |
||
626 |
||
627 |
||
628 |
||
629 |
lemma star_closed_form_bounded_by_rdistinct_list_estimate: |
|
630 |
shows "rsize (rsimp ( RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) |
|
631 |
(star_updates s r0 [[c]]) ) ))) \<le> |
|
632 |
Suc (sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) |
|
633 |
(star_updates s r0 [[c]]) ) {}) ) )" |
|
634 |
by (metis alts_simp_control ) |
|
635 |
||
636 |
||
637 |
||
638 |
||
639 |
lemma star_lambda_form: |
|
640 |
shows "\<forall> r \<in> set (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) ls). |
|
641 |
\<exists>s2. r = RSEQ (rders_simp r0 s2) (RSTAR r0) " |
|
642 |
by (meson ex_map_conv) |
|
643 |
||
644 |
||
645 |
lemma star_lambda_ders: |
|
646 |
shows " \<forall>s. rsize (rders_simp r0 s) \<le> N \<Longrightarrow> |
|
647 |
\<forall>r\<in>set (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) (star_updates s r0 [[c]])). |
|
648 |
rsize r \<le> Suc (N + rsize (RSTAR r0))" |
|
649 |
apply(insert star_lambda_form) |
|
650 |
apply(simp) |
|
651 |
done |
|
652 |
||
653 |
||
654 |
||
655 |
||
656 |
lemma star_control_bounded: |
|
657 |
shows "\<forall>s. rsize (rders_simp r0 s) \<le> N \<Longrightarrow> |
|
658 |
(sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) |
|
659 |
(star_updates s r0 [[c]]) ) {}) ) ) \<le> |
|
660 |
(card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * (Suc (N + rsize (RSTAR r0))) |
|
661 |
" |
|
662 |
apply(subgoal_tac "\<forall>r \<in> set (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) |
|
663 |
(star_updates s r0 [[c]]) ). (rsize r ) \<le> Suc (N + rsize (RSTAR r0))") |
|
664 |
prefer 2 |
|
665 |
using star_lambda_ders apply blast |
|
666 |
using distinct_list_rexp_up_to_certain_size_bouded_by_set_enumerating_up_to_that_size by blast |
|
667 |
||
668 |
||
669 |
lemma star_control_variant: |
|
670 |
assumes "\<forall>s. rsize (rders_simp r0 s) \<le> N" |
|
671 |
shows"Suc |
|
672 |
(sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) |
|
673 |
(star_updates list r0 [[a]])) {}))) |
|
674 |
\<le> (Suc (card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * Suc (N + rsize (RSTAR r0))) " |
|
675 |
apply(subgoal_tac "(sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) |
|
676 |
(star_updates list r0 [[a]])) {}))) |
|
677 |
\<le> ( (card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * Suc (N + rsize (RSTAR r0))) ") |
|
678 |
prefer 2 |
|
679 |
using assms star_control_bounded apply presburger |
|
680 |
by simp |
|
681 |
||
682 |
||
683 |
||
684 |
lemma star_closed_form_bounded: |
|
685 |
shows "\<forall>s. rsize (rders_simp r0 s) \<le> N \<Longrightarrow> |
|
686 |
rsize (rders_simp (RSTAR r0) s) \<le> |
|
687 |
max ( (Suc (card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * (Suc (N + rsize (RSTAR r0))))) (rsize (RSTAR r0))" |
|
688 |
apply(case_tac s) |
|
689 |
apply simp |
|
690 |
apply(subgoal_tac " rsize (rders_simp (RSTAR r0) (a # list)) = |
|
691 |
rsize (rsimp ( RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) (star_updates list r0 [[a]]) ) )))") |
|
692 |
prefer 2 |
|
693 |
using star_closed_form apply presburger |
|
694 |
apply(subgoal_tac "rsize (rsimp ( |
|
695 |
RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) (star_updates list r0 [[a]]) ) ))) |
|
696 |
\<le> Suc (sum_list (map rsize (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) |
|
697 |
(star_updates list r0 [[a]]) ) {}) ) )") |
|
698 |
prefer 2 |
|
699 |
using star_closed_form_bounded_by_rdistinct_list_estimate apply presburger |
|
700 |
apply(subgoal_tac "Suc (sum_list |
|
701 |
(map rsize |
|
702 |
(rdistinct (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0)) (star_updates list r0 [[a]])) {}))) |
|
703 |
\<le> (Suc (card (sizeNregex (Suc (N + rsize (RSTAR r0))))) * Suc (N + rsize (RSTAR r0))) ") |
|
704 |
apply auto[1] |
|
705 |
using star_control_variant by blast |
|
706 |
||
444 | 707 |
lemma seq_estimate_bounded: |
454 | 708 |
assumes "\<forall>s. rsize (rders_simp r1 s) \<le> N1" |
709 |
and "\<forall>s. rsize (rders_simp r2 s) \<le> N2" |
|
444 | 710 |
shows |
454 | 711 |
"Suc (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}))) \<le> |
712 |
Suc (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))" |
|
713 |
proof - |
|
714 |
have a: "sum_list (map rsize (rdistinct (map (rders_simp r2) (vsuf s r1)) {})) \<le> N2 * card (sizeNregex N2)" |
|
715 |
by (metis assms(2) distinct_list_rexp_up_to_certain_size_bouded_by_set_enumerating_up_to_that_size ex_map_conv mult.commute) |
|
444 | 716 |
|
454 | 717 |
have "sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {})) \<le> |
718 |
rsize (RSEQ (rders_simp r1 s) r2) + sum_list (map rsize (rdistinct (map (rders_simp r2) (vsuf s r1)) {}))" |
|
719 |
using triangle_inequality_distinct by blast |
|
720 |
also have "... \<le> rsize (RSEQ (rders_simp r1 s) r2) + N2 * card (sizeNregex N2)" |
|
721 |
by (simp add: a) |
|
722 |
also have "... \<le> Suc (N1 + (rsize r2) + N2 * card (sizeNregex N2))" |
|
723 |
by (simp add: assms(1)) |
|
724 |
finally show ?thesis |
|
725 |
by force |
|
726 |
qed |
|
444 | 727 |
|
452
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
728 |
|
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
729 |
lemma seq_closed_form_bounded2: |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
730 |
assumes "\<forall>s. rsize (rders_simp r1 s) \<le> N1" |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
731 |
and "\<forall>s. rsize (rders_simp r2 s) \<le> N2" |
454 | 732 |
shows "rsize (rders_simp (RSEQ r1 r2) s) |
733 |
\<le> max (Suc (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))) (rsize (RSEQ r1 r2))" |
|
452
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
734 |
proof(cases s) |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
735 |
case Nil |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
736 |
then show "rsize (rders_simp (RSEQ r1 r2) s) |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
737 |
\<le> max (Suc (Suc (N1 + rsize r2) + N2 * card (sizeNregex N2))) (rsize (RSEQ r1 r2))" |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
738 |
by simp |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
739 |
next |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
740 |
case (Cons a list) |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
741 |
then have "rsize (rders_simp (RSEQ r1 r2) s) = |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
742 |
rsize (rsimp (RALTS ((RSEQ (rders_simp r1 s) r2) # (map (rders_simp r2) (vsuf s r1)))))" |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
743 |
using seq_closed_form_variant by (metis list.discI) |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
744 |
also have "... \<le> Suc (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {})))" |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
745 |
using alts_simp_control by blast |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
746 |
also have "... \<le> Suc (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))" |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
747 |
using assms seq_estimate_bounded by blast |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
748 |
ultimately show "rsize (rders_simp (RSEQ r1 r2) s) |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
749 |
\<le> max (Suc (Suc (N1 + rsize r2) + N2 * card (sizeNregex N2))) (rsize (RSEQ r1 r2))" |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
750 |
by auto |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
751 |
qed |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
752 |
|
444 | 753 |
|
452
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
754 |
lemma rders_simp_bounded: |
492 | 755 |
fixes "r" |
452
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
756 |
shows "\<exists>N. \<forall>s. rsize (rders_simp r s) \<le> N" |
444 | 757 |
apply(induct r) |
452
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
758 |
apply(rule_tac x = "Suc 0 " in exI) |
444 | 759 |
using three_easy_cases0 apply force |
760 |
using three_easy_cases1 apply blast |
|
761 |
using three_easy_casesC apply blast |
|
452
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
762 |
apply(erule exE)+ |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
763 |
apply(rule exI) |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
764 |
apply(rule allI) |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
765 |
apply(rule seq_closed_form_bounded2) |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
766 |
apply(assumption) |
4aded96b3923
isarfied one proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
451
diff
changeset
|
767 |
apply(assumption) |
444 | 768 |
apply (metis alts_closed_form_bounded size_list_estimation') |
769 |
using star_closed_form_bounded by blast |
|
770 |
||
492 | 771 |
corollary rders_simp_finiteness: |
772 |
shows "\<forall>r. \<exists>N. \<forall>s. rsize (rders_simp r s) \<le> N" |
|
773 |
using rders_simp_bounded by auto |
|
444 | 774 |
|
775 |
||
776 |
||
777 |
||
778 |
end |