365
+ − 1
+ − 2
theory BitCoded2
+ − 3
imports "Lexer"
+ − 4
begin
+ − 5
+ − 6
section \<open>Bit-Encodings\<close>
+ − 7
+ − 8
datatype bit = Z | S
+ − 9
+ − 10
fun
+ − 11
code :: "val \<Rightarrow> bit list"
+ − 12
where
+ − 13
"code Void = []"
+ − 14
| "code (Char c) = []"
+ − 15
| "code (Left v) = Z # (code v)"
+ − 16
| "code (Right v) = S # (code v)"
+ − 17
| "code (Seq v1 v2) = (code v1) @ (code v2)"
+ − 18
| "code (Stars []) = [S]"
+ − 19
| "code (Stars (v # vs)) = (Z # code v) @ code (Stars vs)"
+ − 20
+ − 21
+ − 22
fun
+ − 23
Stars_add :: "val \<Rightarrow> val \<Rightarrow> val"
+ − 24
where
+ − 25
"Stars_add v (Stars vs) = Stars (v # vs)"
+ − 26
| "Stars_add v _ = Stars [v]"
+ − 27
+ − 28
function
+ − 29
decode' :: "bit list \<Rightarrow> rexp \<Rightarrow> (val * bit list)"
+ − 30
where
+ − 31
"decode' ds ZERO = (Void, [])"
+ − 32
| "decode' ds ONE = (Void, ds)"
+ − 33
| "decode' ds (CHAR d) = (Char d, ds)"
+ − 34
| "decode' [] (ALT r1 r2) = (Void, [])"
+ − 35
| "decode' (Z # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r1 in (Left v, ds'))"
+ − 36
| "decode' (S # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r2 in (Right v, ds'))"
+ − 37
| "decode' ds (SEQ r1 r2) = (let (v1, ds') = decode' ds r1 in
+ − 38
let (v2, ds'') = decode' ds' r2 in (Seq v1 v2, ds''))"
+ − 39
| "decode' [] (STAR r) = (Void, [])"
+ − 40
| "decode' (S # ds) (STAR r) = (Stars [], ds)"
+ − 41
| "decode' (Z # ds) (STAR r) = (let (v, ds') = decode' ds r in
+ − 42
let (vs, ds'') = decode' ds' (STAR r)
+ − 43
in (Stars_add v vs, ds''))"
+ − 44
by pat_completeness auto
+ − 45
+ − 46
lemma decode'_smaller:
+ − 47
assumes "decode'_dom (ds, r)"
+ − 48
shows "length (snd (decode' ds r)) \<le> length ds"
+ − 49
using assms
+ − 50
apply(induct ds r)
+ − 51
apply(auto simp add: decode'.psimps split: prod.split)
+ − 52
using dual_order.trans apply blast
+ − 53
by (meson dual_order.trans le_SucI)
+ − 54
+ − 55
termination "decode'"
+ − 56
apply(relation "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))")
+ − 57
apply(auto dest!: decode'_smaller)
+ − 58
by (metis less_Suc_eq_le snd_conv)
+ − 59
+ − 60
definition
+ − 61
decode :: "bit list \<Rightarrow> rexp \<Rightarrow> val option"
+ − 62
where
+ − 63
"decode ds r \<equiv> (let (v, ds') = decode' ds r
+ − 64
in (if ds' = [] then Some v else None))"
+ − 65
+ − 66
lemma decode'_code_Stars:
+ − 67
assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> (\<forall>x. decode' (code v @ x) r = (v, x)) \<and> flat v \<noteq> []"
+ − 68
shows "decode' (code (Stars vs) @ ds) (STAR r) = (Stars vs, ds)"
+ − 69
using assms
+ − 70
apply(induct vs)
+ − 71
apply(auto)
+ − 72
done
+ − 73
+ − 74
lemma decode'_code:
+ − 75
assumes "\<Turnstile> v : r"
+ − 76
shows "decode' ((code v) @ ds) r = (v, ds)"
+ − 77
using assms
+ − 78
apply(induct v r arbitrary: ds)
+ − 79
apply(auto)
+ − 80
using decode'_code_Stars by blast
+ − 81
+ − 82
lemma decode_code:
+ − 83
assumes "\<Turnstile> v : r"
+ − 84
shows "decode (code v) r = Some v"
+ − 85
using assms unfolding decode_def
+ − 86
by (smt append_Nil2 decode'_code old.prod.case)
+ − 87
+ − 88
+ − 89
section {* Annotated Regular Expressions *}
+ − 90
+ − 91
datatype arexp =
+ − 92
AZERO
+ − 93
| AONE "bit list"
+ − 94
| ACHAR "bit list" char
+ − 95
| ASEQ "bit list" arexp arexp
+ − 96
| AALTs "bit list" "arexp list"
+ − 97
| ASTAR "bit list" arexp
+ − 98
+ − 99
abbreviation
+ − 100
"AALT bs r1 r2 \<equiv> AALTs bs [r1, r2]"
+ − 101
+ − 102
fun asize :: "arexp \<Rightarrow> nat" where
+ − 103
"asize AZERO = 1"
+ − 104
| "asize (AONE cs) = 1"
+ − 105
| "asize (ACHAR cs c) = 1"
+ − 106
| "asize (AALTs cs rs) = Suc (sum_list (map asize rs))"
+ − 107
| "asize (ASEQ cs r1 r2) = Suc (asize r1 + asize r2)"
+ − 108
| "asize (ASTAR cs r) = Suc (asize r)"
+ − 109
+ − 110
fun
+ − 111
erase :: "arexp \<Rightarrow> rexp"
+ − 112
where
+ − 113
"erase AZERO = ZERO"
+ − 114
| "erase (AONE _) = ONE"
+ − 115
| "erase (ACHAR _ c) = CHAR c"
+ − 116
| "erase (AALTs _ []) = ZERO"
+ − 117
| "erase (AALTs _ [r]) = (erase r)"
+ − 118
| "erase (AALTs bs (r#rs)) = ALT (erase r) (erase (AALTs bs rs))"
+ − 119
| "erase (ASEQ _ r1 r2) = SEQ (erase r1) (erase r2)"
+ − 120
| "erase (ASTAR _ r) = STAR (erase r)"
+ − 121
+ − 122
lemma decode_code_erase:
+ − 123
assumes "\<Turnstile> v : (erase a)"
+ − 124
shows "decode (code v) (erase a) = Some v"
+ − 125
using assms
+ − 126
by (simp add: decode_code)
+ − 127
+ − 128
+ − 129
fun nonalt :: "arexp \<Rightarrow> bool"
+ − 130
where
+ − 131
"nonalt (AALTs bs2 rs) = False"
+ − 132
| "nonalt r = True"
+ − 133
+ − 134
+ − 135
fun good :: "arexp \<Rightarrow> bool" where
+ − 136
"good AZERO = False"
+ − 137
| "good (AONE cs) = True"
+ − 138
| "good (ACHAR cs c) = True"
+ − 139
| "good (AALTs cs []) = False"
+ − 140
| "good (AALTs cs [r]) = False"
+ − 141
| "good (AALTs cs (r1#r2#rs)) = (\<forall>r' \<in> set (r1#r2#rs). good r' \<and> nonalt r')"
+ − 142
| "good (ASEQ _ AZERO _) = False"
+ − 143
| "good (ASEQ _ (AONE _) _) = False"
+ − 144
| "good (ASEQ _ _ AZERO) = False"
+ − 145
| "good (ASEQ cs r1 r2) = (good r1 \<and> good r2)"
+ − 146
| "good (ASTAR cs r) = True"
+ − 147
+ − 148
+ − 149
+ − 150
+ − 151
fun fuse :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp" where
+ − 152
"fuse bs AZERO = AZERO"
+ − 153
| "fuse bs (AONE cs) = AONE (bs @ cs)"
+ − 154
| "fuse bs (ACHAR cs c) = ACHAR (bs @ cs) c"
+ − 155
| "fuse bs (AALTs cs rs) = AALTs (bs @ cs) rs"
+ − 156
| "fuse bs (ASEQ cs r1 r2) = ASEQ (bs @ cs) r1 r2"
+ − 157
| "fuse bs (ASTAR cs r) = ASTAR (bs @ cs) r"
+ − 158
+ − 159
lemma fuse_append:
+ − 160
shows "fuse (bs1 @ bs2) r = fuse bs1 (fuse bs2 r)"
+ − 161
apply(induct r)
+ − 162
apply(auto)
+ − 163
done
+ − 164
+ − 165
+ − 166
fun intern :: "rexp \<Rightarrow> arexp" where
+ − 167
"intern ZERO = AZERO"
+ − 168
| "intern ONE = AONE []"
+ − 169
| "intern (CHAR c) = ACHAR [] c"
+ − 170
| "intern (ALT r1 r2) = AALT [] (fuse [Z] (intern r1))
+ − 171
(fuse [S] (intern r2))"
+ − 172
| "intern (SEQ r1 r2) = ASEQ [] (intern r1) (intern r2)"
+ − 173
| "intern (STAR r) = ASTAR [S] (intern r)"
+ − 174
+ − 175
+ − 176
+ − 177
+ − 178
fun retrieve :: "arexp \<Rightarrow> val \<Rightarrow> bit list" where
+ − 179
"retrieve (AONE bs) Void = bs"
+ − 180
| "retrieve (ACHAR bs c) (Char d) = bs"
+ − 181
| "retrieve (AALTs bs [r]) v = bs @ retrieve r v"
+ − 182
| "retrieve (AALTs bs (r#rs)) (Left v) = bs @ retrieve r v"
+ − 183
| "retrieve (AALTs bs (r#rs)) (Right v) = bs @ retrieve (AALTs [] rs) v"
+ − 184
| "retrieve (ASEQ bs r1 r2) (Seq v1 v2) = bs @ retrieve r1 v1 @ retrieve r2 v2"
+ − 185
| "retrieve (ASTAR bs r) (Stars []) = bs @ [S]"
+ − 186
| "retrieve (ASTAR bs r) (Stars (v#vs)) =
+ − 187
bs @ [Z] @ retrieve r v @ retrieve (ASTAR [] r) (Stars vs)"
+ − 188
+ − 189
+ − 190
+ − 191
fun
+ − 192
bnullable :: "arexp \<Rightarrow> bool"
+ − 193
where
+ − 194
"bnullable (AZERO) = False"
+ − 195
| "bnullable (AONE bs) = True"
+ − 196
| "bnullable (ACHAR bs c) = False"
+ − 197
| "bnullable (AALTs bs rs) = (\<exists>r \<in> set rs. bnullable r)"
+ − 198
| "bnullable (ASEQ bs r1 r2) = (bnullable r1 \<and> bnullable r2)"
+ − 199
| "bnullable (ASTAR bs r) = True"
+ − 200
+ − 201
fun
+ − 202
bmkeps :: "arexp \<Rightarrow> bit list"
+ − 203
where
+ − 204
"bmkeps(AONE bs) = bs"
+ − 205
| "bmkeps(ASEQ bs r1 r2) = bs @ (bmkeps r1) @ (bmkeps r2)"
+ − 206
| "bmkeps(AALTs bs [r]) = bs @ (bmkeps r)"
+ − 207
| "bmkeps(AALTs bs (r#rs)) = (if bnullable(r) then bs @ (bmkeps r) else (bmkeps (AALTs bs rs)))"
+ − 208
| "bmkeps(ASTAR bs r) = bs"
+ − 209
+ − 210
+ − 211
fun
+ − 212
bder :: "char \<Rightarrow> arexp \<Rightarrow> arexp"
+ − 213
where
+ − 214
"bder c (AZERO) = AZERO"
+ − 215
| "bder c (AONE bs) = AZERO"
+ − 216
| "bder c (ACHAR bs d) = (if c = d then AONE bs else AZERO)"
+ − 217
| "bder c (AALTs bs rs) = AALTs bs (map (bder c) rs)"
+ − 218
| "bder c (ASEQ bs r1 r2) =
+ − 219
(if bnullable r1
+ − 220
then AALT bs (ASEQ [] (bder c r1) r2) (fuse (bmkeps r1) (bder c r2))
+ − 221
else ASEQ bs (bder c r1) r2)"
+ − 222
| "bder c (ASTAR bs r) = ASEQ (butlast bs) (fuse [Z] (bder c r)) (ASTAR [S] r)"
+ − 223
+ − 224
+ − 225
+ − 226
lemma bder_fuse:
+ − 227
"fuse bs (bder c r) = bder c (fuse bs r)"
+ − 228
apply(induct r arbitrary: bs)
+ − 229
apply(simp_all)
+ − 230
done
+ − 231
+ − 232
+ − 233
fun
+ − 234
bders :: "arexp \<Rightarrow> string \<Rightarrow> arexp"
+ − 235
where
+ − 236
"bders r [] = r"
+ − 237
| "bders r (c#s) = bders (bder c r) s"
+ − 238
+ − 239
lemma bders_append:
+ − 240
"bders r (s1 @ s2) = bders (bders r s1) s2"
+ − 241
apply(induct s1 arbitrary: r s2)
+ − 242
apply(simp_all)
+ − 243
done
+ − 244
+ − 245
lemma bnullable_correctness:
+ − 246
shows "nullable (erase r) = bnullable r"
+ − 247
apply(induct r rule: erase.induct)
+ − 248
apply(simp_all)
+ − 249
done
+ − 250
+ − 251
lemma erase_fuse:
+ − 252
shows "erase (fuse bs r) = erase r"
+ − 253
apply(induct r rule: erase.induct)
+ − 254
apply(simp_all)
+ − 255
done
+ − 256
+ − 257
lemma erase_intern [simp]:
+ − 258
shows "erase (intern r) = r"
+ − 259
apply(induct r)
+ − 260
apply(simp_all add: erase_fuse)
+ − 261
done
+ − 262
+ − 263
lemma erase_bder [simp]:
+ − 264
shows "erase (bder a r) = der a (erase r)"
+ − 265
apply(induct r rule: erase.induct)
+ − 266
apply(simp_all add: erase_fuse bnullable_correctness)
+ − 267
done
+ − 268
+ − 269
lemma erase_bders [simp]:
+ − 270
shows "erase (bders r s) = ders s (erase r)"
+ − 271
apply(induct s arbitrary: r )
+ − 272
apply(simp_all)
+ − 273
done
+ − 274
+ − 275
lemma retrieve_encode_STARS:
+ − 276
assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> code v = retrieve (intern r) v"
+ − 277
shows "code (Stars vs) = retrieve (ASTAR [] (intern r)) (Stars vs)"
+ − 278
using assms
+ − 279
apply(induct vs)
+ − 280
apply(simp_all)
+ − 281
done
+ − 282
+ − 283
lemma retrieve_fuse2:
+ − 284
assumes "\<Turnstile> v : (erase r)"
+ − 285
shows "retrieve (fuse bs r) v = bs @ retrieve r v"
+ − 286
using assms
+ − 287
apply(induct r arbitrary: v bs)
+ − 288
apply(auto elim: Prf_elims)[4]
+ − 289
defer
+ − 290
using retrieve_encode_STARS
+ − 291
apply(auto elim!: Prf_elims)[1]
+ − 292
apply(case_tac vs)
+ − 293
apply(simp)
+ − 294
apply(simp)
+ − 295
(* AALTs case *)
+ − 296
apply(simp)
+ − 297
apply(case_tac x2a)
+ − 298
apply(simp)
+ − 299
apply(auto elim!: Prf_elims)[1]
+ − 300
apply(simp)
+ − 301
apply(case_tac list)
+ − 302
apply(simp)
+ − 303
apply(auto)
+ − 304
apply(auto elim!: Prf_elims)[1]
+ − 305
done
+ − 306
+ − 307
lemma retrieve_fuse:
+ − 308
assumes "\<Turnstile> v : r"
+ − 309
shows "retrieve (fuse bs (intern r)) v = bs @ retrieve (intern r) v"
+ − 310
using assms
+ − 311
by (simp_all add: retrieve_fuse2)
+ − 312
+ − 313
+ − 314
lemma r:
+ − 315
assumes "bnullable (AALTs bs (a # rs))"
+ − 316
shows "bnullable a \<or> (\<not> bnullable a \<and> bnullable (AALTs bs rs))"
+ − 317
using assms
+ − 318
apply(induct rs)
+ − 319
apply(auto)
+ − 320
done
+ − 321
+ − 322
lemma r0:
+ − 323
assumes "bnullable a"
+ − 324
shows "bmkeps (AALTs bs (a # rs)) = bs @ (bmkeps a)"
+ − 325
using assms
+ − 326
by (metis bmkeps.simps(3) bmkeps.simps(4) list.exhaust)
+ − 327
+ − 328
lemma r1:
+ − 329
assumes "\<not> bnullable a" "bnullable (AALTs bs rs)"
+ − 330
shows "bmkeps (AALTs bs (a # rs)) = bmkeps (AALTs bs rs)"
+ − 331
using assms
+ − 332
apply(induct rs)
+ − 333
apply(auto)
+ − 334
done
+ − 335
+ − 336
lemma r2:
+ − 337
assumes "x \<in> set rs" "bnullable x"
+ − 338
shows "bnullable (AALTs bs rs)"
+ − 339
using assms
+ − 340
apply(induct rs)
+ − 341
apply(auto)
+ − 342
done
+ − 343
+ − 344
lemma r3:
+ − 345
assumes "\<not> bnullable r"
+ − 346
" \<exists> x \<in> set rs. bnullable x"
+ − 347
shows "retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs))) =
+ − 348
retrieve (AALTs bs (r # rs)) (mkeps (erase (AALTs bs (r # rs))))"
+ − 349
using assms
+ − 350
apply(induct rs arbitrary: r bs)
+ − 351
apply(auto)[1]
+ − 352
apply(auto)
+ − 353
using bnullable_correctness apply blast
+ − 354
apply(auto simp add: bnullable_correctness mkeps_nullable retrieve_fuse2)
+ − 355
apply(subst retrieve_fuse2[symmetric])
+ − 356
apply (smt bnullable.simps(4) bnullable_correctness erase.simps(5) erase.simps(6) insert_iff list.exhaust list.set(2) mkeps.simps(3) mkeps_nullable)
+ − 357
apply(simp)
+ − 358
apply(case_tac "bnullable a")
+ − 359
apply (smt append_Nil2 bnullable.simps(4) bnullable_correctness erase.simps(5) erase.simps(6) fuse.simps(4) insert_iff list.exhaust list.set(2) mkeps.simps(3) mkeps_nullable retrieve_fuse2)
+ − 360
apply(drule_tac x="a" in meta_spec)
+ − 361
apply(drule_tac x="bs" in meta_spec)
+ − 362
apply(drule meta_mp)
+ − 363
apply(simp)
+ − 364
apply(drule meta_mp)
+ − 365
apply(auto)
+ − 366
apply(subst retrieve_fuse2[symmetric])
+ − 367
apply(case_tac rs)
+ − 368
apply(simp)
+ − 369
apply(auto)[1]
+ − 370
apply (simp add: bnullable_correctness)
+ − 371
apply (metis append_Nil2 bnullable_correctness erase_fuse fuse.simps(4) list.set_intros(1) mkeps.simps(3) mkeps_nullable nullable.simps(4) r2)
+ − 372
apply (simp add: bnullable_correctness)
+ − 373
apply (metis append_Nil2 bnullable_correctness erase.simps(6) erase_fuse fuse.simps(4) list.set_intros(2) mkeps.simps(3) mkeps_nullable r2)
+ − 374
apply(simp)
+ − 375
done
+ − 376
+ − 377
+ − 378
lemma t:
+ − 379
assumes "\<forall>r \<in> set rs. nullable (erase r) \<longrightarrow> bmkeps r = retrieve r (mkeps (erase r))"
+ − 380
"nullable (erase (AALTs bs rs))"
+ − 381
shows " bmkeps (AALTs bs rs) = retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs)))"
+ − 382
using assms
+ − 383
apply(induct rs arbitrary: bs)
+ − 384
apply(simp)
+ − 385
apply(auto simp add: bnullable_correctness)
+ − 386
apply(case_tac rs)
+ − 387
apply(auto simp add: bnullable_correctness)[2]
+ − 388
apply(subst r1)
+ − 389
apply(simp)
+ − 390
apply(rule r2)
+ − 391
apply(assumption)
+ − 392
apply(simp)
+ − 393
apply(drule_tac x="bs" in meta_spec)
+ − 394
apply(drule meta_mp)
+ − 395
apply(auto)[1]
+ − 396
prefer 2
+ − 397
apply(case_tac "bnullable a")
+ − 398
apply(subst r0)
+ − 399
apply blast
+ − 400
apply(subgoal_tac "nullable (erase a)")
+ − 401
prefer 2
+ − 402
using bnullable_correctness apply blast
+ − 403
apply (metis (no_types, lifting) erase.simps(5) erase.simps(6) list.exhaust mkeps.simps(3) retrieve.simps(3) retrieve.simps(4))
+ − 404
apply(subst r1)
+ − 405
apply(simp)
+ − 406
using r2 apply blast
+ − 407
apply(drule_tac x="bs" in meta_spec)
+ − 408
apply(drule meta_mp)
+ − 409
apply(auto)[1]
+ − 410
apply(simp)
+ − 411
using r3 apply blast
+ − 412
apply(auto)
+ − 413
using r3 by blast
+ − 414
+ − 415
+ − 416
lemma asize0:
+ − 417
shows "0 < asize r"
+ − 418
apply(induct r)
+ − 419
apply(auto)
+ − 420
done
+ − 421
+ − 422
lemma asize_fuse:
+ − 423
shows "asize (fuse bs r) = asize r"
+ − 424
apply(induct r)
+ − 425
apply(auto)
+ − 426
done
+ − 427
+ − 428
lemma TESTTEST:
+ − 429
shows "bder c (intern r) = intern (der c r)"
+ − 430
apply(induct r)
+ − 431
apply(simp)
+ − 432
apply(simp)
+ − 433
apply(simp)
+ − 434
prefer 2
+ − 435
apply(simp)
+ − 436
apply (simp add: bder_fuse[symmetric])
+ − 437
prefer 3
+ − 438
apply(simp only: intern.simps)
+ − 439
apply(simp only: der.simps)
+ − 440
apply(simp only: intern.simps)
+ − 441
apply(simp only: bder.simps)
+ − 442
apply(simp)
+ − 443
apply(simp only: intern.simps)
+ − 444
prefer 2
+ − 445
apply(simp)
+ − 446
prefer 2
+ − 447
apply(simp)
+ − 448
apply(auto)
+ − 449
+ − 450
+ − 451
fun nonnested :: "arexp \<Rightarrow> bool"
+ − 452
where
+ − 453
"nonnested (AALTs bs2 []) = True"
+ − 454
| "nonnested (AALTs bs2 ((AALTs bs1 rs1) # rs2)) = False"
+ − 455
| "nonnested (AALTs bs2 (r # rs2)) = nonnested (AALTs bs2 rs2)"
+ − 456
| "nonnested r = True"
+ − 457
+ − 458
+ − 459
+ − 460
fun distinctBy :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b set \<Rightarrow> 'a list"
+ − 461
where
+ − 462
"distinctBy [] f acc = []"
+ − 463
| "distinctBy (x#xs) f acc =
+ − 464
(if (f x) \<in> acc then distinctBy xs f acc
+ − 465
else x # (distinctBy xs f ({f x} \<union> acc)))"
+ − 466
+ − 467
fun flts :: "arexp list \<Rightarrow> arexp list"
+ − 468
where
+ − 469
"flts [] = []"
+ − 470
| "flts (AZERO # rs) = flts rs"
+ − 471
| "flts ((AALTs bs rs1) # rs) = (map (fuse bs) rs1) @ flts rs"
+ − 472
| "flts (r1 # rs) = r1 # flts rs"
+ − 473
+ − 474
+ − 475
fun spill :: "arexp list \<Rightarrow> arexp list"
+ − 476
where
+ − 477
"spill [] = []"
+ − 478
| "spill ((AALTs bs rs1) # rs) = (map (fuse bs) rs1) @ spill rs"
+ − 479
| "spill (r1 # rs) = r1 # spill rs"
+ − 480
+ − 481
lemma spill_Cons:
+ − 482
shows "spill (r # rs1) = spill [r] @ spill rs1"
+ − 483
apply(induct r arbitrary: rs1)
+ − 484
apply(auto)
+ − 485
done
+ − 486
+ − 487
lemma spill_append:
+ − 488
shows "spill (rs1 @ rs2) = spill rs1 @ spill rs2"
+ − 489
apply(induct rs1 arbitrary: rs2)
+ − 490
apply(auto)
+ − 491
by (metis append.assoc spill_Cons)
+ − 492
+ − 493
fun bsimp_ASEQ :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp \<Rightarrow> arexp"
+ − 494
where
+ − 495
"bsimp_ASEQ _ AZERO _ = AZERO"
+ − 496
| "bsimp_ASEQ _ _ AZERO = AZERO"
+ − 497
| "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2"
+ − 498
| "bsimp_ASEQ bs1 r1 r2 = ASEQ bs1 r1 r2"
+ − 499
+ − 500
+ − 501
fun bsimp_AALTs :: "bit list \<Rightarrow> arexp list \<Rightarrow> arexp"
+ − 502
where
+ − 503
"bsimp_AALTs _ [] = AZERO"
+ − 504
| "bsimp_AALTs bs1 [r] = fuse bs1 r"
+ − 505
| "bsimp_AALTs bs1 rs = AALTs bs1 rs"
+ − 506
+ − 507
+ − 508
fun bsimp :: "arexp \<Rightarrow> arexp"
+ − 509
where
+ − 510
"bsimp (ASEQ bs1 r1 r2) = bsimp_ASEQ bs1 (bsimp r1) (bsimp r2)"
+ − 511
| "bsimp (AALTs bs1 rs) = bsimp_AALTs bs1 (flts (map bsimp rs))"
+ − 512
| "bsimp r = r"
+ − 513
+ − 514
+ − 515
inductive contains2 :: "arexp \<Rightarrow> bit list \<Rightarrow> bool" ("_ >>2 _" [51, 50] 50)
+ − 516
where
+ − 517
"AONE bs >>2 bs"
+ − 518
| "ACHAR bs c >>2 bs"
+ − 519
| "\<lbrakk>a1 >>2 bs1; a2 >>2 bs2\<rbrakk> \<Longrightarrow> ASEQ bs a1 a2 >>2 bs @ bs1 @ bs2"
+ − 520
| "r >>2 bs1 \<Longrightarrow> AALTs bs (r#rs) >>2 bs @ bs1"
+ − 521
| "AALTs bs rs >>2 bs @ bs1 \<Longrightarrow> AALTs bs (r#rs) >>2 bs @ bs1"
+ − 522
| "ASTAR bs r >>2 bs @ [S]"
+ − 523
| "\<lbrakk>r >>2 bs1; ASTAR [] r >>2 bs2\<rbrakk> \<Longrightarrow> ASTAR bs r >>2 bs @ Z # bs1 @ bs2"
+ − 524
| "r >>2 bs \<Longrightarrow> (bsimp r) >>2 bs"
+ − 525
+ − 526
+ − 527
inductive contains :: "arexp \<Rightarrow> bit list \<Rightarrow> bool" ("_ >> _" [51, 50] 50)
+ − 528
where
+ − 529
"AONE bs >> bs"
+ − 530
| "ACHAR bs c >> bs"
+ − 531
| "\<lbrakk>a1 >> bs1; a2 >> bs2\<rbrakk> \<Longrightarrow> ASEQ bs a1 a2 >> bs @ bs1 @ bs2"
+ − 532
| "r >> bs1 \<Longrightarrow> AALTs bs (r#rs) >> bs @ bs1"
+ − 533
| "AALTs bs rs >> bs @ bs1 \<Longrightarrow> AALTs bs (r#rs) >> bs @ bs1"
+ − 534
| "ASTAR bs r >> bs @ [S]"
+ − 535
| "\<lbrakk>r >> bs1; ASTAR [] r >> bs2\<rbrakk> \<Longrightarrow> ASTAR bs r >> bs @ Z # bs1 @ bs2"
+ − 536
+ − 537
+ − 538
+ − 539
lemma contains0:
+ − 540
assumes "a >> bs"
+ − 541
shows "(fuse bs1 a) >> bs1 @ bs"
+ − 542
using assms
+ − 543
apply(induct arbitrary: bs1)
+ − 544
apply(auto intro: contains.intros)
+ − 545
apply (metis append.assoc contains.intros(3))
+ − 546
apply (metis append.assoc contains.intros(4))
+ − 547
apply (metis append.assoc contains.intros(5))
+ − 548
apply (metis append.assoc contains.intros(6))
+ − 549
apply (metis append_assoc contains.intros(7))
+ − 550
done
+ − 551
+ − 552
lemma contains1:
+ − 553
assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> intern r >> code v"
+ − 554
shows "ASTAR [] (intern r) >> code (Stars vs)"
+ − 555
using assms
+ − 556
apply(induct vs)
+ − 557
apply(simp)
+ − 558
using contains.simps apply blast
+ − 559
apply(simp)
+ − 560
apply(subst (2) append_Nil[symmetric])
+ − 561
apply(rule contains.intros)
+ − 562
apply(auto)
+ − 563
done
+ − 564
+ − 565
+ − 566
+ − 567
+ − 568
+ − 569
lemma contains2:
+ − 570
assumes "\<Turnstile> v : r"
+ − 571
shows "(intern r) >> code v"
+ − 572
using assms
+ − 573
apply(induct)
+ − 574
prefer 4
+ − 575
apply(simp)
+ − 576
apply(rule contains.intros)
+ − 577
prefer 4
+ − 578
apply(simp)
+ − 579
apply(rule contains.intros)
+ − 580
apply(simp)
+ − 581
apply(subst (3) append_Nil[symmetric])
+ − 582
apply(rule contains.intros)
+ − 583
apply(simp)
+ − 584
apply(simp)
+ − 585
apply(simp)
+ − 586
apply(subst (9) append_Nil[symmetric])
+ − 587
apply(rule contains.intros)
+ − 588
apply (metis append_Cons append_self_conv2 contains0)
+ − 589
apply(simp)
+ − 590
apply(subst (9) append_Nil[symmetric])
+ − 591
apply(rule contains.intros)
+ − 592
back
+ − 593
apply(rule contains.intros)
+ − 594
apply(drule_tac ?bs1.0="[S]" in contains0)
+ − 595
apply(simp)
+ − 596
apply(simp)
+ − 597
apply(case_tac vs)
+ − 598
apply(simp)
+ − 599
apply (metis append_Nil contains.intros(6))
+ − 600
using contains1 by blast
+ − 601
+ − 602
lemma qq1:
+ − 603
assumes "\<exists>r \<in> set rs. bnullable r"
+ − 604
shows "bmkeps (AALTs bs (rs @ rs1)) = bmkeps (AALTs bs rs)"
+ − 605
using assms
+ − 606
apply(induct rs arbitrary: rs1 bs)
+ − 607
apply(simp)
+ − 608
apply(simp)
+ − 609
by (metis Nil_is_append_conv bmkeps.simps(4) neq_Nil_conv r0 split_list_last)
+ − 610
+ − 611
lemma qq2:
+ − 612
assumes "\<forall>r \<in> set rs. \<not> bnullable r" "\<exists>r \<in> set rs1. bnullable r"
+ − 613
shows "bmkeps (AALTs bs (rs @ rs1)) = bmkeps (AALTs bs rs1)"
+ − 614
using assms
+ − 615
apply(induct rs arbitrary: rs1 bs)
+ − 616
apply(simp)
+ − 617
apply(simp)
+ − 618
by (metis append_assoc in_set_conv_decomp r1 r2)
+ − 619
+ − 620
lemma qq2a:
+ − 621
assumes "\<not> bnullable r" "\<exists>r \<in> set rs1. bnullable r"
+ − 622
shows "bmkeps (AALTs bs (r # rs1)) = bmkeps (AALTs bs rs1)"
+ − 623
using assms
+ − 624
by (simp add: r1)
+ − 625
+ − 626
lemma qq3:
+ − 627
shows "bnullable (AALTs bs rs) = (\<exists>r \<in> set rs. bnullable r)"
+ − 628
apply(induct rs arbitrary: bs)
+ − 629
apply(simp)
+ − 630
apply(simp)
+ − 631
done
+ − 632
+ − 633
lemma qq4:
+ − 634
assumes "bnullable (AALTs bs rs)"
+ − 635
shows "bmkeps (AALTs bs rs) = bs @ bmkeps (AALTs [] rs)"
+ − 636
by (metis append_Nil2 assms bmkeps_retrieve bnullable_correctness erase_fuse fuse.simps(4) mkeps_nullable retrieve_fuse2)
+ − 637
+ − 638
+ − 639
lemma contains3a:
+ − 640
assumes "AALTs bs lst >> bs @ bs1"
+ − 641
shows "AALTs bs (a # lst) >> bs @ bs1"
+ − 642
using assms
+ − 643
apply -
+ − 644
by (simp add: contains.intros(5))
+ − 645
+ − 646
+ − 647
lemma contains3b:
+ − 648
assumes "a >> bs1"
+ − 649
shows "AALTs bs (a # lst) >> bs @ bs1"
+ − 650
using assms
+ − 651
apply -
+ − 652
apply(rule contains.intros)
+ − 653
apply(simp)
+ − 654
done
+ − 655
+ − 656
+ − 657
lemma contains3:
+ − 658
assumes "\<And>x. \<lbrakk>x \<in> set rs; bnullable x\<rbrakk> \<Longrightarrow> x >> bmkeps x" "x \<in> set rs" "bnullable x"
+ − 659
shows "AALTs bs rs >> bmkeps (AALTs bs rs)"
+ − 660
using assms
+ − 661
apply(induct rs arbitrary: bs x)
+ − 662
apply simp
+ − 663
by (metis contains.intros(4) contains.intros(5) list.set_intros(1) list.set_intros(2) qq3 qq4 r r0 r1)
+ − 664
+ − 665
lemma cont1:
+ − 666
assumes "\<And>v. \<Turnstile> v : erase r \<Longrightarrow> r >> retrieve r v"
+ − 667
"\<forall>v\<in>set vs. \<Turnstile> v : erase r \<and> flat v \<noteq> []"
+ − 668
shows "ASTAR bs r >> retrieve (ASTAR bs r) (Stars vs)"
+ − 669
using assms
+ − 670
apply(induct vs arbitrary: bs r)
+ − 671
apply(simp)
+ − 672
using contains.intros(6) apply auto[1]
+ − 673
by (simp add: contains.intros(7))
+ − 674
+ − 675
lemma contains4:
+ − 676
assumes "bnullable a"
+ − 677
shows "a >> bmkeps a"
+ − 678
using assms
+ − 679
apply(induct a rule: bnullable.induct)
+ − 680
apply(auto intro: contains.intros)
+ − 681
using contains3 by blast
+ − 682
+ − 683
lemma contains5:
+ − 684
assumes "\<Turnstile> v : r"
+ − 685
shows "(intern r) >> retrieve (intern r) v"
+ − 686
using contains2[OF assms] retrieve_code[OF assms]
+ − 687
by (simp)
+ − 688
+ − 689
+ − 690
lemma contains6:
+ − 691
assumes "\<Turnstile> v : (erase r)"
+ − 692
shows "r >> retrieve r v"
+ − 693
using assms
+ − 694
apply(induct r arbitrary: v rule: erase.induct)
+ − 695
apply(auto)[1]
+ − 696
using Prf_elims(1) apply blast
+ − 697
using Prf_elims(4) contains.intros(1) apply force
+ − 698
using Prf_elims(5) contains.intros(2) apply force
+ − 699
apply(auto)[1]
+ − 700
using Prf_elims(1) apply blast
+ − 701
apply(auto)[1]
+ − 702
using contains3b contains3a apply blast
+ − 703
prefer 2
+ − 704
apply(auto)[1]
+ − 705
apply (metis Prf_elims(2) contains.intros(3) retrieve.simps(6))
+ − 706
prefer 2
+ − 707
apply(auto)[1]
+ − 708
apply (metis Prf_elims(6) cont1)
+ − 709
apply(simp)
+ − 710
apply(erule Prf_elims)
+ − 711
apply(auto)
+ − 712
apply (simp add: contains3b)
+ − 713
using retrieve_fuse2 contains3b contains3a
+ − 714
apply(subst retrieve_fuse2[symmetric])
+ − 715
apply (metis append_Nil2 erase_fuse fuse.simps(4))
+ − 716
apply(simp)
+ − 717
by (metis append_Nil2 erase_fuse fuse.simps(4))
+ − 718
+ − 719
lemma contains7:
+ − 720
assumes "\<Turnstile> v : der c (erase r)"
+ − 721
shows "(bder c r) >> retrieve r (injval (erase r) c v)"
+ − 722
using bder_retrieve[OF assms(1)] retrieve_code[OF assms(1)]
+ − 723
by (metis assms contains6 erase_bder)
+ − 724
+ − 725
+ − 726
lemma contains7a:
+ − 727
assumes "\<Turnstile> v : der c (erase r)"
+ − 728
shows "r >> retrieve r (injval (erase r) c v)"
+ − 729
using assms
+ − 730
apply -
+ − 731
apply(drule Prf_injval)
+ − 732
apply(drule contains6)
+ − 733
apply(simp)
+ − 734
done
+ − 735
+ − 736
lemma contains7b:
+ − 737
assumes "\<Turnstile> v : ders s (erase r)"
+ − 738
shows "(bders r s) >> retrieve r (flex (erase r) id s v)"
+ − 739
using assms
+ − 740
apply(induct s arbitrary: r v)
+ − 741
apply(simp)
+ − 742
apply (simp add: contains6)
+ − 743
apply(simp add: bders_append flex_append ders_append)
+ − 744
apply(drule_tac x="bder a r" in meta_spec)
+ − 745
apply(drule meta_spec)
+ − 746
apply(drule meta_mp)
+ − 747
apply(simp)
+ − 748
apply(simp)
+ − 749
apply(subst (asm) bder_retrieve)
+ − 750
defer
+ − 751
apply (simp add: flex_injval)
+ − 752
by (simp add: Prf_flex)
+ − 753
+ − 754
lemma contains7_iff:
+ − 755
assumes "\<Turnstile> v : der c (erase r)"
+ − 756
shows "(bder c r) >> retrieve r (injval (erase r) c v) \<longleftrightarrow>
+ − 757
r >> retrieve r (injval (erase r) c v)"
+ − 758
by (simp add: assms contains7 contains7a)
+ − 759
+ − 760
lemma contains8_iff:
+ − 761
assumes "\<Turnstile> v : ders s (erase r)"
+ − 762
shows "(bders r s) >> retrieve r (flex (erase r) id s v) \<longleftrightarrow>
+ − 763
r >> retrieve r (flex (erase r) id s v)"
+ − 764
using Prf_flex assms contains6 contains7b by blast
+ − 765
+ − 766
+ − 767
+ − 768
+ − 769
fun
+ − 770
bders_simp :: "arexp \<Rightarrow> string \<Rightarrow> arexp"
+ − 771
where
+ − 772
"bders_simp r [] = r"
+ − 773
| "bders_simp r (c # s) = bders_simp (bsimp (bder c r)) s"
+ − 774
+ − 775
definition blexer_simp where
+ − 776
"blexer_simp r s \<equiv> if bnullable (bders_simp (intern r) s) then
+ − 777
decode (bmkeps (bders_simp (intern r) s)) r else None"
+ − 778
+ − 779
+ − 780
+ − 781
+ − 782
+ − 783
lemma bders_simp_append:
+ − 784
shows "bders_simp r (s1 @ s2) = bders_simp (bders_simp r s1) s2"
+ − 785
apply(induct s1 arbitrary: r s2)
+ − 786
apply(simp)
+ − 787
apply(simp)
+ − 788
done
+ − 789
+ − 790
lemma bsimp_ASEQ_size:
+ − 791
shows "asize (bsimp_ASEQ bs r1 r2) \<le> Suc (asize r1 + asize r2)"
+ − 792
apply(induct bs r1 r2 rule: bsimp_ASEQ.induct)
+ − 793
apply(auto)
+ − 794
done
+ − 795
+ − 796
+ − 797
+ − 798
lemma flts_size:
+ − 799
shows "sum_list (map asize (flts rs)) \<le> sum_list (map asize rs)"
+ − 800
apply(induct rs rule: flts.induct)
+ − 801
apply(simp_all)
+ − 802
by (simp add: asize_fuse comp_def)
+ − 803
+ − 804
+ − 805
lemma bsimp_AALTs_size:
+ − 806
shows "asize (bsimp_AALTs bs rs) \<le> Suc (sum_list (map asize rs))"
+ − 807
apply(induct rs rule: bsimp_AALTs.induct)
+ − 808
apply(auto simp add: asize_fuse)
+ − 809
done
+ − 810
+ − 811
+ − 812
lemma bsimp_size:
+ − 813
shows "asize (bsimp r) \<le> asize r"
+ − 814
apply(induct r)
+ − 815
apply(simp_all)
+ − 816
apply (meson Suc_le_mono add_mono_thms_linordered_semiring(1) bsimp_ASEQ_size le_trans)
+ − 817
apply(rule le_trans)
+ − 818
apply(rule bsimp_AALTs_size)
+ − 819
apply(simp)
+ − 820
apply(rule le_trans)
+ − 821
apply(rule flts_size)
+ − 822
by (simp add: sum_list_mono)
+ − 823
+ − 824
lemma bsimp_asize0:
+ − 825
shows "(\<Sum>x\<leftarrow>rs. asize (bsimp x)) \<le> sum_list (map asize rs)"
+ − 826
apply(induct rs)
+ − 827
apply(auto)
+ − 828
by (simp add: add_mono bsimp_size)
+ − 829
+ − 830
lemma bsimp_AALTs_size2:
+ − 831
assumes "\<forall>r \<in> set rs. nonalt r"
+ − 832
shows "asize (bsimp_AALTs bs rs) \<ge> sum_list (map asize rs)"
+ − 833
using assms
+ − 834
apply(induct rs rule: bsimp_AALTs.induct)
+ − 835
apply(simp_all add: asize_fuse)
+ − 836
done
+ − 837
+ − 838
+ − 839
lemma qq:
+ − 840
shows "map (asize \<circ> fuse bs) rs = map asize rs"
+ − 841
apply(induct rs)
+ − 842
apply(auto simp add: asize_fuse)
+ − 843
done
+ − 844
+ − 845
lemma flts_size2:
+ − 846
assumes "\<exists>bs rs'. AALTs bs rs' \<in> set rs"
+ − 847
shows "sum_list (map asize (flts rs)) < sum_list (map asize rs)"
+ − 848
using assms
+ − 849
apply(induct rs)
+ − 850
apply(auto simp add: qq)
+ − 851
apply (simp add: flts_size less_Suc_eq_le)
+ − 852
apply(case_tac a)
+ − 853
apply(auto simp add: qq)
+ − 854
prefer 2
+ − 855
apply (simp add: flts_size le_imp_less_Suc)
+ − 856
using less_Suc_eq by auto
+ − 857
+ − 858
lemma bsimp_AALTs_size3:
+ − 859
assumes "\<exists>r \<in> set (map bsimp rs). \<not>nonalt r"
+ − 860
shows "asize (bsimp (AALTs bs rs)) < asize (AALTs bs rs)"
+ − 861
using assms flts_size2
+ − 862
apply -
+ − 863
apply(clarify)
+ − 864
apply(simp)
+ − 865
apply(drule_tac x="map bsimp rs" in meta_spec)
+ − 866
apply(drule meta_mp)
+ − 867
apply (metis list.set_map nonalt.elims(3))
+ − 868
apply(simp)
+ − 869
apply(rule order_class.order.strict_trans1)
+ − 870
apply(rule bsimp_AALTs_size)
+ − 871
apply(simp)
+ − 872
by (smt Suc_leI bsimp_asize0 comp_def le_imp_less_Suc le_trans map_eq_conv not_less_eq)
+ − 873
+ − 874
+ − 875
+ − 876
+ − 877
lemma L_bsimp_ASEQ:
+ − 878
"L (SEQ (erase r1) (erase r2)) = L (erase (bsimp_ASEQ bs r1 r2))"
+ − 879
apply(induct bs r1 r2 rule: bsimp_ASEQ.induct)
+ − 880
apply(simp_all)
+ − 881
by (metis erase_fuse fuse.simps(4))
+ − 882
+ − 883
lemma L_bsimp_AALTs:
+ − 884
"L (erase (AALTs bs rs)) = L (erase (bsimp_AALTs bs rs))"
+ − 885
apply(induct bs rs rule: bsimp_AALTs.induct)
+ − 886
apply(simp_all add: erase_fuse)
+ − 887
done
+ − 888
+ − 889
lemma L_erase_AALTs:
+ − 890
shows "L (erase (AALTs bs rs)) = \<Union> (L ` erase ` (set rs))"
+ − 891
apply(induct rs)
+ − 892
apply(simp)
+ − 893
apply(simp)
+ − 894
apply(case_tac rs)
+ − 895
apply(simp)
+ − 896
apply(simp)
+ − 897
done
+ − 898
+ − 899
lemma L_erase_flts:
+ − 900
shows "\<Union> (L ` erase ` (set (flts rs))) = \<Union> (L ` erase ` (set rs))"
+ − 901
apply(induct rs rule: flts.induct)
+ − 902
apply(simp_all)
+ − 903
apply(auto)
+ − 904
using L_erase_AALTs erase_fuse apply auto[1]
+ − 905
by (simp add: L_erase_AALTs erase_fuse)
+ − 906
+ − 907
+ − 908
lemma L_bsimp_erase:
+ − 909
shows "L (erase r) = L (erase (bsimp r))"
+ − 910
apply(induct r)
+ − 911
apply(simp)
+ − 912
apply(simp)
+ − 913
apply(simp)
+ − 914
apply(auto simp add: Sequ_def)[1]
+ − 915
apply(subst L_bsimp_ASEQ[symmetric])
+ − 916
apply(auto simp add: Sequ_def)[1]
+ − 917
apply(subst (asm) L_bsimp_ASEQ[symmetric])
+ − 918
apply(auto simp add: Sequ_def)[1]
+ − 919
apply(simp)
+ − 920
apply(subst L_bsimp_AALTs[symmetric])
+ − 921
defer
+ − 922
apply(simp)
+ − 923
apply(subst (2)L_erase_AALTs)
+ − 924
apply(subst L_erase_flts)
+ − 925
apply(auto)
+ − 926
apply (simp add: L_erase_AALTs)
+ − 927
using L_erase_AALTs by blast
+ − 928
+ − 929
lemma bsimp_ASEQ0:
+ − 930
shows "bsimp_ASEQ bs r1 AZERO = AZERO"
+ − 931
apply(induct r1)
+ − 932
apply(auto)
+ − 933
done
+ − 934
+ − 935
+ − 936
+ − 937
lemma bsimp_ASEQ1:
+ − 938
assumes "r1 \<noteq> AZERO" "r2 \<noteq> AZERO" "\<forall>bs. r1 \<noteq> AONE bs"
+ − 939
shows "bsimp_ASEQ bs r1 r2 = ASEQ bs r1 r2"
+ − 940
using assms
+ − 941
apply(induct bs r1 r2 rule: bsimp_ASEQ.induct)
+ − 942
apply(auto)
+ − 943
done
+ − 944
+ − 945
lemma bsimp_ASEQ2:
+ − 946
shows "bsimp_ASEQ bs (AONE bs1) r2 = fuse (bs @ bs1) r2"
+ − 947
apply(induct r2)
+ − 948
apply(auto)
+ − 949
done
+ − 950
+ − 951
+ − 952
lemma L_bders_simp:
+ − 953
shows "L (erase (bders_simp r s)) = L (erase (bders r s))"
+ − 954
apply(induct s arbitrary: r rule: rev_induct)
+ − 955
apply(simp)
+ − 956
apply(simp)
+ − 957
apply(simp add: ders_append)
+ − 958
apply(simp add: bders_simp_append)
+ − 959
apply(simp add: L_bsimp_erase[symmetric])
+ − 960
by (simp add: der_correctness)
+ − 961
+ − 962
lemma b1:
+ − 963
"bsimp_ASEQ bs1 (AONE bs) r = fuse (bs1 @ bs) r"
+ − 964
apply(induct r)
+ − 965
apply(auto)
+ − 966
done
+ − 967
+ − 968
lemma b2:
+ − 969
assumes "bnullable r"
+ − 970
shows "bmkeps (fuse bs r) = bs @ bmkeps r"
+ − 971
by (simp add: assms bmkeps_retrieve bnullable_correctness erase_fuse mkeps_nullable retrieve_fuse2)
+ − 972
+ − 973
lemma b3:
+ − 974
shows "bnullable r = bnullable (bsimp r)"
+ − 975
using L_bsimp_erase bnullable_correctness nullable_correctness by auto
+ − 976
+ − 977
+ − 978
lemma b4:
+ − 979
shows "bnullable (bders_simp r s) = bnullable (bders r s)"
+ − 980
by (metis L_bders_simp bnullable_correctness lexer.simps(1) lexer_correct_None option.distinct(1))
+ − 981
+ − 982
lemma q1:
+ − 983
assumes "\<forall>r \<in> set rs. bmkeps(bsimp r) = bmkeps r"
+ − 984
shows "map (\<lambda>r. bmkeps(bsimp r)) rs = map bmkeps rs"
+ − 985
using assms
+ − 986
apply(induct rs)
+ − 987
apply(simp)
+ − 988
apply(simp)
+ − 989
done
+ − 990
+ − 991
lemma q3:
+ − 992
assumes "\<exists>r \<in> set rs. bnullable r"
+ − 993
shows "bmkeps (AALTs bs rs) = bmkeps (bsimp_AALTs bs rs)"
+ − 994
using assms
+ − 995
apply(induct bs rs rule: bsimp_AALTs.induct)
+ − 996
apply(simp)
+ − 997
apply(simp)
+ − 998
apply (simp add: b2)
+ − 999
apply(simp)
+ − 1000
done
+ − 1001
+ − 1002
+ − 1003
lemma fuse_empty:
+ − 1004
shows "fuse [] r = r"
+ − 1005
apply(induct r)
+ − 1006
apply(auto)
+ − 1007
done
+ − 1008
+ − 1009
lemma flts_fuse:
+ − 1010
shows "map (fuse bs) (flts rs) = flts (map (fuse bs) rs)"
+ − 1011
apply(induct rs arbitrary: bs rule: flts.induct)
+ − 1012
apply(auto simp add: fuse_append)
+ − 1013
done
+ − 1014
+ − 1015
lemma bsimp_ASEQ_fuse:
+ − 1016
shows "fuse bs1 (bsimp_ASEQ bs2 r1 r2) = bsimp_ASEQ (bs1 @ bs2) r1 r2"
+ − 1017
apply(induct r1 r2 arbitrary: bs1 bs2 rule: bsimp_ASEQ.induct)
+ − 1018
apply(auto)
+ − 1019
done
+ − 1020
+ − 1021
lemma bsimp_AALTs_fuse:
+ − 1022
assumes "\<forall>r \<in> set rs. fuse bs1 (fuse bs2 r) = fuse (bs1 @ bs2) r"
+ − 1023
shows "fuse bs1 (bsimp_AALTs bs2 rs) = bsimp_AALTs (bs1 @ bs2) rs"
+ − 1024
using assms
+ − 1025
apply(induct bs2 rs arbitrary: bs1 rule: bsimp_AALTs.induct)
+ − 1026
apply(auto)
+ − 1027
done
+ − 1028
+ − 1029
+ − 1030
+ − 1031
lemma bsimp_fuse:
+ − 1032
shows "fuse bs (bsimp r) = bsimp (fuse bs r)"
+ − 1033
apply(induct r arbitrary: bs)
+ − 1034
apply(simp)
+ − 1035
apply(simp)
+ − 1036
apply(simp)
+ − 1037
prefer 3
+ − 1038
apply(simp)
+ − 1039
apply(simp)
+ − 1040
apply (simp add: bsimp_ASEQ_fuse)
+ − 1041
apply(simp)
+ − 1042
by (simp add: bsimp_AALTs_fuse fuse_append)
+ − 1043
+ − 1044
lemma bsimp_fuse_AALTs:
+ − 1045
shows "fuse bs (bsimp (AALTs [] rs)) = bsimp (AALTs bs rs)"
+ − 1046
apply(subst bsimp_fuse)
+ − 1047
apply(simp)
+ − 1048
done
+ − 1049
+ − 1050
lemma bsimp_fuse_AALTs2:
+ − 1051
shows "fuse bs (bsimp_AALTs [] rs) = bsimp_AALTs bs rs"
+ − 1052
using bsimp_AALTs_fuse fuse_append by auto
+ − 1053
+ − 1054
+ − 1055
lemma bsimp_ASEQ_idem:
+ − 1056
assumes "bsimp (bsimp r1) = bsimp r1" "bsimp (bsimp r2) = bsimp r2"
+ − 1057
shows "bsimp (bsimp_ASEQ x1 (bsimp r1) (bsimp r2)) = bsimp_ASEQ x1 (bsimp r1) (bsimp r2)"
+ − 1058
using assms
+ − 1059
apply(case_tac "bsimp r1 = AZERO")
+ − 1060
apply(simp)
+ − 1061
apply(case_tac "bsimp r2 = AZERO")
+ − 1062
apply(simp)
+ − 1063
apply (metis bnullable.elims(2) bnullable.elims(3) bsimp.simps(3) bsimp_ASEQ.simps(2) bsimp_ASEQ.simps(3) bsimp_ASEQ.simps(4) bsimp_ASEQ.simps(5) bsimp_ASEQ.simps(6))
+ − 1064
apply(case_tac "\<exists>bs. bsimp r1 = AONE bs")
+ − 1065
apply(auto)[1]
+ − 1066
apply(subst bsimp_ASEQ2)
+ − 1067
apply(subst bsimp_ASEQ2)
+ − 1068
apply (metis assms(2) bsimp_fuse)
+ − 1069
apply(subst bsimp_ASEQ1)
+ − 1070
apply(auto)
+ − 1071
done
+ − 1072
+ − 1073
+ − 1074
+ − 1075
lemma k0:
+ − 1076
shows "flts (r # rs1) = flts [r] @ flts rs1"
+ − 1077
apply(induct r arbitrary: rs1)
+ − 1078
apply(auto)
+ − 1079
done
+ − 1080
+ − 1081
lemma k00:
+ − 1082
shows "flts (rs1 @ rs2) = flts rs1 @ flts rs2"
+ − 1083
apply(induct rs1 arbitrary: rs2)
+ − 1084
apply(auto)
+ − 1085
by (metis append.assoc k0)
+ − 1086
+ − 1087
lemma k0a:
+ − 1088
shows "flts [AALTs bs rs] = map (fuse bs) rs"
+ − 1089
apply(simp)
+ − 1090
done
+ − 1091
+ − 1092
+ − 1093
lemma k0b:
+ − 1094
assumes "nonalt r" "r \<noteq> AZERO"
+ − 1095
shows "flts [r] = [r]"
+ − 1096
using assms
+ − 1097
apply(case_tac r)
+ − 1098
apply(simp_all)
+ − 1099
done
+ − 1100
+ − 1101
lemma nn1:
+ − 1102
assumes "nonnested (AALTs bs rs)"
+ − 1103
shows "\<nexists>bs1 rs1. flts rs = [AALTs bs1 rs1]"
+ − 1104
using assms
+ − 1105
apply(induct rs rule: flts.induct)
+ − 1106
apply(auto)
+ − 1107
done
+ − 1108
+ − 1109
lemma nn1q:
+ − 1110
assumes "nonnested (AALTs bs rs)"
+ − 1111
shows "\<nexists>bs1 rs1. AALTs bs1 rs1 \<in> set (flts rs)"
+ − 1112
using assms
+ − 1113
apply(induct rs rule: flts.induct)
+ − 1114
apply(auto)
+ − 1115
done
+ − 1116
+ − 1117
lemma nn1qq:
+ − 1118
assumes "nonnested (AALTs bs rs)"
+ − 1119
shows "\<nexists>bs1 rs1. AALTs bs1 rs1 \<in> set rs"
+ − 1120
using assms
+ − 1121
apply(induct rs rule: flts.induct)
+ − 1122
apply(auto)
+ − 1123
done
+ − 1124
+ − 1125
lemma nn10:
+ − 1126
assumes "nonnested (AALTs cs rs)"
+ − 1127
shows "nonnested (AALTs (bs @ cs) rs)"
+ − 1128
using assms
+ − 1129
apply(induct rs arbitrary: cs bs)
+ − 1130
apply(simp_all)
+ − 1131
apply(case_tac a)
+ − 1132
apply(simp_all)
+ − 1133
done
+ − 1134
+ − 1135
lemma nn11a:
+ − 1136
assumes "nonalt r"
+ − 1137
shows "nonalt (fuse bs r)"
+ − 1138
using assms
+ − 1139
apply(induct r)
+ − 1140
apply(auto)
+ − 1141
done
+ − 1142
+ − 1143
+ − 1144
lemma nn1a:
+ − 1145
assumes "nonnested r"
+ − 1146
shows "nonnested (fuse bs r)"
+ − 1147
using assms
+ − 1148
apply(induct bs r arbitrary: rule: fuse.induct)
+ − 1149
apply(simp_all add: nn10)
+ − 1150
done
+ − 1151
+ − 1152
lemma n0:
+ − 1153
shows "nonnested (AALTs bs rs) \<longleftrightarrow> (\<forall>r \<in> set rs. nonalt r)"
+ − 1154
apply(induct rs arbitrary: bs)
+ − 1155
apply(auto)
+ − 1156
apply (metis list.set_intros(1) nn1qq nonalt.elims(3))
+ − 1157
apply (metis list.set_intros(2) nn1qq nonalt.elims(3))
+ − 1158
by (metis nonalt.elims(2) nonnested.simps(3) nonnested.simps(4) nonnested.simps(5) nonnested.simps(6) nonnested.simps(7))
+ − 1159
+ − 1160
+ − 1161
+ − 1162
+ − 1163
lemma nn1c:
+ − 1164
assumes "\<forall>r \<in> set rs. nonnested r"
+ − 1165
shows "\<forall>r \<in> set (flts rs). nonalt r"
+ − 1166
using assms
+ − 1167
apply(induct rs rule: flts.induct)
+ − 1168
apply(auto)
+ − 1169
apply(rule nn11a)
+ − 1170
by (metis nn1qq nonalt.elims(3))
+ − 1171
+ − 1172
lemma nn1bb:
+ − 1173
assumes "\<forall>r \<in> set rs. nonalt r"
+ − 1174
shows "nonnested (bsimp_AALTs bs rs)"
+ − 1175
using assms
+ − 1176
apply(induct bs rs rule: bsimp_AALTs.induct)
+ − 1177
apply(auto)
+ − 1178
apply (metis nn11a nonalt.simps(1) nonnested.elims(3))
+ − 1179
using n0 by auto
+ − 1180
+ − 1181
lemma nn1b:
+ − 1182
shows "nonnested (bsimp r)"
+ − 1183
apply(induct r)
+ − 1184
apply(simp_all)
+ − 1185
apply(case_tac "bsimp r1 = AZERO")
+ − 1186
apply(simp)
+ − 1187
apply(case_tac "bsimp r2 = AZERO")
+ − 1188
apply(simp)
+ − 1189
apply(subst bsimp_ASEQ0)
+ − 1190
apply(simp)
+ − 1191
apply(case_tac "\<exists>bs. bsimp r1 = AONE bs")
+ − 1192
apply(auto)[1]
+ − 1193
apply(subst bsimp_ASEQ2)
+ − 1194
apply (simp add: nn1a)
+ − 1195
apply(subst bsimp_ASEQ1)
+ − 1196
apply(auto)
+ − 1197
apply(rule nn1bb)
+ − 1198
apply(auto)
+ − 1199
by (metis (mono_tags, hide_lams) imageE nn1c set_map)
+ − 1200
+ − 1201
lemma nn1d:
+ − 1202
assumes "bsimp r = AALTs bs rs"
+ − 1203
shows "\<forall>r1 \<in> set rs. \<forall> bs. r1 \<noteq> AALTs bs rs2"
+ − 1204
using nn1b assms
+ − 1205
by (metis nn1qq)
+ − 1206
+ − 1207
lemma nn_flts:
+ − 1208
assumes "nonnested (AALTs bs rs)"
+ − 1209
shows "\<forall>r \<in> set (flts rs). nonalt r"
+ − 1210
using assms
+ − 1211
apply(induct rs arbitrary: bs rule: flts.induct)
+ − 1212
apply(auto)
+ − 1213
done
+ − 1214
+ − 1215
+ − 1216
+ − 1217
lemma rt:
+ − 1218
shows "sum_list (map asize (flts (map bsimp rs))) \<le> sum_list (map asize rs)"
+ − 1219
apply(induct rs)
+ − 1220
apply(simp)
+ − 1221
apply(simp)
+ − 1222
apply(subst k0)
+ − 1223
apply(simp)
+ − 1224
by (smt add_le_cancel_right add_mono bsimp_size flts.simps(1) flts_size k0 le_iff_add list.simps(9) map_append sum_list.Cons sum_list.append trans_le_add1)
+ − 1225
+ − 1226
lemma bsimp_AALTs_qq:
+ − 1227
assumes "1 < length rs"
+ − 1228
shows "bsimp_AALTs bs rs = AALTs bs rs"
+ − 1229
using assms
+ − 1230
apply(case_tac rs)
+ − 1231
apply(simp)
+ − 1232
apply(case_tac list)
+ − 1233
apply(simp_all)
+ − 1234
done
+ − 1235
+ − 1236
+ − 1237
lemma bsimp_AALTs1:
+ − 1238
assumes "nonalt r"
+ − 1239
shows "bsimp_AALTs bs (flts [r]) = fuse bs r"
+ − 1240
using assms
+ − 1241
apply(case_tac r)
+ − 1242
apply(simp_all)
+ − 1243
done
+ − 1244
+ − 1245
lemma bbbbs:
+ − 1246
assumes "good r" "r = AALTs bs1 rs"
+ − 1247
shows "bsimp_AALTs bs (flts [r]) = AALTs bs (map (fuse bs1) rs)"
+ − 1248
using assms
+ − 1249
by (metis (no_types, lifting) Nil_is_map_conv append.left_neutral append_butlast_last_id bsimp_AALTs.elims butlast.simps(2) good.simps(4) good.simps(5) k0a map_butlast)
+ − 1250
+ − 1251
lemma bbbbs1:
+ − 1252
shows "nonalt r \<or> (\<exists>bs rs. r = AALTs bs rs)"
+ − 1253
using nonalt.elims(3) by auto
+ − 1254
+ − 1255
+ − 1256
lemma good_fuse:
+ − 1257
shows "good (fuse bs r) = good r"
+ − 1258
apply(induct r arbitrary: bs)
+ − 1259
apply(auto)
+ − 1260
apply(case_tac r1)
+ − 1261
apply(simp_all)
+ − 1262
apply(case_tac r2)
+ − 1263
apply(simp_all)
+ − 1264
apply(case_tac r2)
+ − 1265
apply(simp_all)
+ − 1266
apply(case_tac r2)
+ − 1267
apply(simp_all)
+ − 1268
apply(case_tac r2)
+ − 1269
apply(simp_all)
+ − 1270
apply(case_tac r1)
+ − 1271
apply(simp_all)
+ − 1272
apply(case_tac r2)
+ − 1273
apply(simp_all)
+ − 1274
apply(case_tac r2)
+ − 1275
apply(simp_all)
+ − 1276
apply(case_tac r2)
+ − 1277
apply(simp_all)
+ − 1278
apply(case_tac r2)
+ − 1279
apply(simp_all)
+ − 1280
apply(case_tac x2a)
+ − 1281
apply(simp_all)
+ − 1282
apply(case_tac list)
+ − 1283
apply(simp_all)
+ − 1284
apply(case_tac x2a)
+ − 1285
apply(simp_all)
+ − 1286
apply(case_tac list)
+ − 1287
apply(simp_all)
+ − 1288
done
+ − 1289
+ − 1290
lemma good0:
+ − 1291
assumes "rs \<noteq> Nil" "\<forall>r \<in> set rs. nonalt r"
+ − 1292
shows "good (bsimp_AALTs bs rs) \<longleftrightarrow> (\<forall>r \<in> set rs. good r)"
+ − 1293
using assms
+ − 1294
apply(induct bs rs rule: bsimp_AALTs.induct)
+ − 1295
apply(auto simp add: good_fuse)
+ − 1296
done
+ − 1297
+ − 1298
lemma good0a:
+ − 1299
assumes "flts (map bsimp rs) \<noteq> Nil" "\<forall>r \<in> set (flts (map bsimp rs)). nonalt r"
+ − 1300
shows "good (bsimp (AALTs bs rs)) \<longleftrightarrow> (\<forall>r \<in> set (flts (map bsimp rs)). good r)"
+ − 1301
using assms
+ − 1302
apply(simp)
+ − 1303
apply(auto)
+ − 1304
apply(subst (asm) good0)
+ − 1305
apply(simp)
+ − 1306
apply(auto)
+ − 1307
apply(subst good0)
+ − 1308
apply(simp)
+ − 1309
apply(auto)
+ − 1310
done
+ − 1311
+ − 1312
lemma flts0:
+ − 1313
assumes "r \<noteq> AZERO" "nonalt r"
+ − 1314
shows "flts [r] \<noteq> []"
+ − 1315
using assms
+ − 1316
apply(induct r)
+ − 1317
apply(simp_all)
+ − 1318
done
+ − 1319
+ − 1320
lemma flts1:
+ − 1321
assumes "good r"
+ − 1322
shows "flts [r] \<noteq> []"
+ − 1323
using assms
+ − 1324
apply(induct r)
+ − 1325
apply(simp_all)
+ − 1326
apply(case_tac x2a)
+ − 1327
apply(simp)
+ − 1328
apply(simp)
+ − 1329
done
+ − 1330
+ − 1331
lemma flts2:
+ − 1332
assumes "good r"
+ − 1333
shows "\<forall>r' \<in> set (flts [r]). good r' \<and> nonalt r'"
+ − 1334
using assms
+ − 1335
apply(induct r)
+ − 1336
apply(simp)
+ − 1337
apply(simp)
+ − 1338
apply(simp)
+ − 1339
prefer 2
+ − 1340
apply(simp)
+ − 1341
apply(auto)[1]
+ − 1342
apply (metis bsimp_AALTs.elims good.simps(4) good.simps(5) good.simps(6) good_fuse)
+ − 1343
apply (metis bsimp_AALTs.elims good.simps(4) good.simps(5) good.simps(6) nn11a)
+ − 1344
apply fastforce
+ − 1345
apply(simp)
+ − 1346
done
+ − 1347
+ − 1348
+ − 1349
lemma flts3:
+ − 1350
assumes "\<forall>r \<in> set rs. good r \<or> r = AZERO"
+ − 1351
shows "\<forall>r \<in> set (flts rs). good r"
+ − 1352
using assms
+ − 1353
apply(induct rs arbitrary: rule: flts.induct)
+ − 1354
apply(simp_all)
+ − 1355
by (metis UnE flts2 k0a set_map)
+ − 1356
+ − 1357
lemma flts3b:
+ − 1358
assumes "\<exists>r\<in>set rs. good r"
+ − 1359
shows "flts rs \<noteq> []"
+ − 1360
using assms
+ − 1361
apply(induct rs arbitrary: rule: flts.induct)
+ − 1362
apply(simp)
+ − 1363
apply(simp)
+ − 1364
apply(simp)
+ − 1365
apply(auto)
+ − 1366
done
+ − 1367
+ − 1368
lemma flts4:
+ − 1369
assumes "bsimp_AALTs bs (flts rs) = AZERO"
+ − 1370
shows "\<forall>r \<in> set rs. \<not> good r"
+ − 1371
using assms
+ − 1372
apply(induct rs arbitrary: bs rule: flts.induct)
+ − 1373
apply(auto)
+ − 1374
defer
+ − 1375
apply (metis (no_types, lifting) Nil_is_append_conv append_self_conv2 bsimp_AALTs.elims butlast.simps(2) butlast_append flts3b nonalt.simps(1) nonalt.simps(2))
+ − 1376
apply (metis arexp.distinct(7) bsimp_AALTs.elims flts2 good.simps(1) good.simps(2) good0 k0b list.distinct(1) list.inject nonalt.simps(3))
+ − 1377
apply (metis arexp.distinct(3) arexp.distinct(7) bsimp_AALTs.elims fuse.simps(3) list.distinct(1) list.inject)
+ − 1378
apply (metis arexp.distinct(7) bsimp_AALTs.elims good.simps(1) good_fuse list.distinct(1) list.inject)
+ − 1379
apply (metis arexp.distinct(7) bsimp_AALTs.elims list.distinct(1) list.inject)
+ − 1380
apply (metis arexp.distinct(7) bsimp_AALTs.elims flts2 good.simps(1) good.simps(33) good0 k0b list.distinct(1) list.inject nonalt.simps(6))
+ − 1381
by (metis (no_types, lifting) Nil_is_append_conv append_Nil2 arexp.distinct(7) bsimp_AALTs.elims butlast.simps(2) butlast_append flts1 flts2 good.simps(1) good0 k0a)
+ − 1382
+ − 1383
+ − 1384
lemma flts_nil:
+ − 1385
assumes "\<forall>y. asize y < Suc (sum_list (map asize rs)) \<longrightarrow>
+ − 1386
good (bsimp y) \<or> bsimp y = AZERO"
+ − 1387
and "\<forall>r\<in>set rs. \<not> good (bsimp r)"
+ − 1388
shows "flts (map bsimp rs) = []"
+ − 1389
using assms
+ − 1390
apply(induct rs)
+ − 1391
apply(simp)
+ − 1392
apply(simp)
+ − 1393
apply(subst k0)
+ − 1394
apply(simp)
+ − 1395
by force
+ − 1396
+ − 1397
lemma flts_nil2:
+ − 1398
assumes "\<forall>y. asize y < Suc (sum_list (map asize rs)) \<longrightarrow>
+ − 1399
good (bsimp y) \<or> bsimp y = AZERO"
+ − 1400
and "bsimp_AALTs bs (flts (map bsimp rs)) = AZERO"
+ − 1401
shows "flts (map bsimp rs) = []"
+ − 1402
using assms
+ − 1403
apply(induct rs arbitrary: bs)
+ − 1404
apply(simp)
+ − 1405
apply(simp)
+ − 1406
apply(subst k0)
+ − 1407
apply(simp)
+ − 1408
apply(subst (asm) k0)
+ − 1409
apply(auto)
+ − 1410
apply (metis flts.simps(1) flts.simps(2) flts4 k0 less_add_Suc1 list.set_intros(1))
+ − 1411
by (metis flts.simps(2) flts4 k0 less_add_Suc1 list.set_intros(1))
+ − 1412
+ − 1413
+ − 1414
+ − 1415
lemma good_SEQ:
+ − 1416
assumes "r1 \<noteq> AZERO" "r2 \<noteq> AZERO" "\<forall>bs. r1 \<noteq> AONE bs"
+ − 1417
shows "good (ASEQ bs r1 r2) = (good r1 \<and> good r2)"
+ − 1418
using assms
+ − 1419
apply(case_tac r1)
+ − 1420
apply(simp_all)
+ − 1421
apply(case_tac r2)
+ − 1422
apply(simp_all)
+ − 1423
apply(case_tac r2)
+ − 1424
apply(simp_all)
+ − 1425
apply(case_tac r2)
+ − 1426
apply(simp_all)
+ − 1427
apply(case_tac r2)
+ − 1428
apply(simp_all)
+ − 1429
done
+ − 1430
+ − 1431
lemma good1:
+ − 1432
shows "good (bsimp a) \<or> bsimp a = AZERO"
+ − 1433
apply(induct a taking: asize rule: measure_induct)
+ − 1434
apply(case_tac x)
+ − 1435
apply(simp)
+ − 1436
apply(simp)
+ − 1437
apply(simp)
+ − 1438
prefer 3
+ − 1439
apply(simp)
+ − 1440
prefer 2
+ − 1441
(* AALTs case *)
+ − 1442
apply(simp only:)
+ − 1443
apply(case_tac "x52")
+ − 1444
apply(simp)
+ − 1445
thm good0a
+ − 1446
(* AALTs list at least one - case *)
+ − 1447
apply(simp only: )
+ − 1448
apply(frule_tac x="a" in spec)
+ − 1449
apply(drule mp)
+ − 1450
apply(simp)
+ − 1451
(* either first element is good, or AZERO *)
+ − 1452
apply(erule disjE)
+ − 1453
prefer 2
+ − 1454
apply(simp)
+ − 1455
(* in the AZERO case, the size is smaller *)
+ − 1456
apply(drule_tac x="AALTs x51 list" in spec)
+ − 1457
apply(drule mp)
+ − 1458
apply(simp add: asize0)
+ − 1459
apply(subst (asm) bsimp.simps)
+ − 1460
apply(subst (asm) bsimp.simps)
+ − 1461
apply(assumption)
+ − 1462
(* in the good case *)
+ − 1463
apply(frule_tac x="AALTs x51 list" in spec)
+ − 1464
apply(drule mp)
+ − 1465
apply(simp add: asize0)
+ − 1466
apply(erule disjE)
+ − 1467
apply(rule disjI1)
+ − 1468
apply(simp add: good0)
+ − 1469
apply(subst good0)
+ − 1470
apply (metis Nil_is_append_conv flts1 k0)
+ − 1471
apply (metis ex_map_conv list.simps(9) nn1b nn1c)
+ − 1472
apply(simp)
+ − 1473
apply(subst k0)
+ − 1474
apply(simp)
+ − 1475
apply(auto)[1]
+ − 1476
using flts2 apply blast
+ − 1477
apply(subst (asm) good0)
+ − 1478
prefer 3
+ − 1479
apply(auto)[1]
+ − 1480
apply auto[1]
+ − 1481
apply (metis ex_map_conv nn1b nn1c)
+ − 1482
(* in the AZERO case *)
+ − 1483
apply(simp)
+ − 1484
apply(frule_tac x="a" in spec)
+ − 1485
apply(drule mp)
+ − 1486
apply(simp)
+ − 1487
apply(erule disjE)
+ − 1488
apply(rule disjI1)
+ − 1489
apply(subst good0)
+ − 1490
apply(subst k0)
+ − 1491
using flts1 apply blast
+ − 1492
apply(auto)[1]
+ − 1493
apply (metis (no_types, hide_lams) ex_map_conv list.simps(9) nn1b nn1c)
+ − 1494
apply(auto)[1]
+ − 1495
apply(subst (asm) k0)
+ − 1496
apply(auto)[1]
+ − 1497
using flts2 apply blast
+ − 1498
apply(frule_tac x="AALTs x51 list" in spec)
+ − 1499
apply(drule mp)
+ − 1500
apply(simp add: asize0)
+ − 1501
apply(erule disjE)
+ − 1502
apply(simp)
+ − 1503
apply(simp)
+ − 1504
apply (metis add.left_commute flts_nil2 less_add_Suc1 less_imp_Suc_add list.distinct(1) list.set_cases nat.inject)
+ − 1505
apply(subst (2) k0)
+ − 1506
apply(simp)
+ − 1507
(* SEQ case *)
+ − 1508
apply(simp)
+ − 1509
apply(case_tac "bsimp x42 = AZERO")
+ − 1510
apply(simp)
+ − 1511
apply(case_tac "bsimp x43 = AZERO")
+ − 1512
apply(simp)
+ − 1513
apply(subst (2) bsimp_ASEQ0)
+ − 1514
apply(simp)
+ − 1515
apply(case_tac "\<exists>bs. bsimp x42 = AONE bs")
+ − 1516
apply(auto)[1]
+ − 1517
apply(subst bsimp_ASEQ2)
+ − 1518
using good_fuse apply force
+ − 1519
apply(subst bsimp_ASEQ1)
+ − 1520
apply(auto)
+ − 1521
apply(subst good_SEQ)
+ − 1522
apply(simp)
+ − 1523
apply(simp)
+ − 1524
apply(simp)
+ − 1525
using less_add_Suc1 less_add_Suc2 by blast
+ − 1526
+ − 1527
lemma good1a:
+ − 1528
assumes "L(erase a) \<noteq> {}"
+ − 1529
shows "good (bsimp a)"
+ − 1530
using good1 assms
+ − 1531
using L_bsimp_erase by force
+ − 1532
+ − 1533
+ − 1534
+ − 1535
lemma flts_append:
+ − 1536
"flts (xs1 @ xs2) = flts xs1 @ flts xs2"
+ − 1537
apply(induct xs1 arbitrary: xs2 rule: rev_induct)
+ − 1538
apply(auto)
+ − 1539
apply(case_tac xs)
+ − 1540
apply(auto)
+ − 1541
apply(case_tac x)
+ − 1542
apply(auto)
+ − 1543
apply(case_tac x)
+ − 1544
apply(auto)
+ − 1545
done
+ − 1546
+ − 1547
lemma g1:
+ − 1548
assumes "good (bsimp_AALTs bs rs)"
+ − 1549
shows "bsimp_AALTs bs rs = AALTs bs rs \<or> (\<exists>r. rs = [r] \<and> bsimp_AALTs bs [r] = fuse bs r)"
+ − 1550
using assms
+ − 1551
apply(induct rs arbitrary: bs)
+ − 1552
apply(simp)
+ − 1553
apply(case_tac rs)
+ − 1554
apply(simp only:)
+ − 1555
apply(simp)
+ − 1556
apply(case_tac list)
+ − 1557
apply(simp)
+ − 1558
by simp
+ − 1559
+ − 1560
lemma flts_0:
+ − 1561
assumes "nonnested (AALTs bs rs)"
+ − 1562
shows "\<forall>r \<in> set (flts rs). r \<noteq> AZERO"
+ − 1563
using assms
+ − 1564
apply(induct rs arbitrary: bs rule: flts.induct)
+ − 1565
apply(simp)
+ − 1566
apply(simp)
+ − 1567
defer
+ − 1568
apply(simp)
+ − 1569
apply(simp)
+ − 1570
apply(simp)
+ − 1571
apply(simp)
+ − 1572
apply(rule ballI)
+ − 1573
apply(simp)
+ − 1574
done
+ − 1575
+ − 1576
lemma flts_0a:
+ − 1577
assumes "nonnested (AALTs bs rs)"
+ − 1578
shows "AZERO \<notin> set (flts rs)"
+ − 1579
using assms
+ − 1580
using flts_0 by blast
+ − 1581
+ − 1582
lemma qqq1:
+ − 1583
shows "AZERO \<notin> set (flts (map bsimp rs))"
+ − 1584
by (metis ex_map_conv flts3 good.simps(1) good1)
+ − 1585
+ − 1586
+ − 1587
fun nonazero :: "arexp \<Rightarrow> bool"
+ − 1588
where
+ − 1589
"nonazero AZERO = False"
+ − 1590
| "nonazero r = True"
+ − 1591
+ − 1592
lemma flts_concat:
+ − 1593
shows "flts rs = concat (map (\<lambda>r. flts [r]) rs)"
+ − 1594
apply(induct rs)
+ − 1595
apply(auto)
+ − 1596
apply(subst k0)
+ − 1597
apply(simp)
+ − 1598
done
+ − 1599
+ − 1600
lemma flts_single1:
+ − 1601
assumes "nonalt r" "nonazero r"
+ − 1602
shows "flts [r] = [r]"
+ − 1603
using assms
+ − 1604
apply(induct r)
+ − 1605
apply(auto)
+ − 1606
done
+ − 1607
+ − 1608
lemma flts_qq:
+ − 1609
assumes "\<forall>y. asize y < Suc (sum_list (map asize rs)) \<longrightarrow> good y \<longrightarrow> bsimp y = y"
+ − 1610
"\<forall>r'\<in>set rs. good r' \<and> nonalt r'"
+ − 1611
shows "flts (map bsimp rs) = rs"
+ − 1612
using assms
+ − 1613
apply(induct rs)
+ − 1614
apply(simp)
+ − 1615
apply(simp)
+ − 1616
apply(subst k0)
+ − 1617
apply(subgoal_tac "flts [bsimp a] = [a]")
+ − 1618
prefer 2
+ − 1619
apply(drule_tac x="a" in spec)
+ − 1620
apply(drule mp)
+ − 1621
apply(simp)
+ − 1622
apply(auto)[1]
+ − 1623
using good.simps(1) k0b apply blast
+ − 1624
apply(auto)[1]
+ − 1625
done
+ − 1626
+ − 1627
lemma test:
+ − 1628
assumes "good r"
+ − 1629
shows "bsimp r = r"
+ − 1630
using assms
+ − 1631
apply(induct r taking: "asize" rule: measure_induct)
+ − 1632
apply(erule good.elims)
+ − 1633
apply(simp_all)
+ − 1634
apply(subst k0)
+ − 1635
apply(subst (2) k0)
+ − 1636
apply(subst flts_qq)
+ − 1637
apply(auto)[1]
+ − 1638
apply(auto)[1]
+ − 1639
apply (metis append_Cons append_Nil bsimp_AALTs.simps(3) good.simps(1) k0b)
+ − 1640
apply force+
+ − 1641
apply (metis (no_types, lifting) add_Suc add_Suc_right asize.simps(5) bsimp.simps(1) bsimp_ASEQ.simps(19) less_add_Suc1 less_add_Suc2)
+ − 1642
apply (metis add_Suc add_Suc_right arexp.distinct(5) arexp.distinct(7) asize.simps(4) asize.simps(5) bsimp.simps(1) bsimp.simps(2) bsimp_ASEQ1 good.simps(21) good.simps(8) less_add_Suc1 less_add_Suc2)
+ − 1643
apply force+
+ − 1644
apply (metis (no_types, lifting) add_Suc add_Suc_right arexp.distinct(5) arexp.distinct(7) asize.simps(4) asize.simps(5) bsimp.simps(1) bsimp.simps(2) bsimp_ASEQ1 good.simps(25) good.simps(8) less_add_Suc1 less_add_Suc2)
+ − 1645
apply (metis add_Suc add_Suc_right arexp.distinct(7) asize.simps(4) bsimp.simps(2) bsimp_ASEQ1 good.simps(26) good.simps(8) less_add_Suc1 less_add_Suc2)
+ − 1646
apply force+
+ − 1647
done
+ − 1648
+ − 1649
lemma test2:
+ − 1650
assumes "good r"
+ − 1651
shows "bsimp r = r"
+ − 1652
using assms
+ − 1653
apply(induct r taking: "asize" rule: measure_induct)
+ − 1654
apply(case_tac x)
+ − 1655
apply(simp_all)
+ − 1656
defer
+ − 1657
(* AALT case *)
+ − 1658
apply(subgoal_tac "1 < length x52")
+ − 1659
prefer 2
+ − 1660
apply(case_tac x52)
+ − 1661
apply(simp)
+ − 1662
apply(simp)
+ − 1663
apply(case_tac list)
+ − 1664
apply(simp)
+ − 1665
apply(simp)
+ − 1666
apply(subst bsimp_AALTs_qq)
+ − 1667
prefer 2
+ − 1668
apply(subst flts_qq)
+ − 1669
apply(auto)[1]
+ − 1670
apply(auto)[1]
+ − 1671
apply(case_tac x52)
+ − 1672
apply(simp)
+ − 1673
apply(simp)
+ − 1674
apply(case_tac list)
+ − 1675
apply(simp)
+ − 1676
apply(simp)
+ − 1677
apply(auto)[1]
+ − 1678
apply (metis (no_types, lifting) bsimp_AALTs.elims good.simps(6) length_Cons length_pos_if_in_set list.size(3) nat_neq_iff)
+ − 1679
apply(simp)
+ − 1680
apply(case_tac x52)
+ − 1681
apply(simp)
+ − 1682
apply(simp)
+ − 1683
apply(case_tac list)
+ − 1684
apply(simp)
+ − 1685
apply(simp)
+ − 1686
apply(subst k0)
+ − 1687
apply(simp)
+ − 1688
apply(subst (2) k0)
+ − 1689
apply(simp)
+ − 1690
apply (simp add: Suc_lessI flts1 one_is_add)
+ − 1691
(* SEQ case *)
+ − 1692
apply(case_tac "bsimp x42 = AZERO")
+ − 1693
apply simp
+ − 1694
apply (metis asize.elims good.simps(10) good.simps(11) good.simps(12) good.simps(2) good.simps(7) good.simps(9) good_SEQ less_add_Suc1)
+ − 1695
apply(case_tac "\<exists>bs'. bsimp x42 = AONE bs'")
+ − 1696
apply(auto)[1]
+ − 1697
defer
+ − 1698
apply(case_tac "bsimp x43 = AZERO")
+ − 1699
apply(simp)
+ − 1700
apply (metis bsimp.elims bsimp.simps(3) good.simps(10) good.simps(11) good.simps(12) good.simps(8) good.simps(9) good_SEQ less_add_Suc2)
+ − 1701
apply(auto)
+ − 1702
apply (subst bsimp_ASEQ1)
+ − 1703
apply(auto)[3]
+ − 1704
apply(auto)[1]
+ − 1705
apply (metis bsimp.simps(3) good.simps(2) good_SEQ less_add_Suc1)
+ − 1706
apply (metis bsimp.simps(3) good.simps(2) good_SEQ less_add_Suc1 less_add_Suc2)
+ − 1707
apply (subst bsimp_ASEQ2)
+ − 1708
apply(drule_tac x="x42" in spec)
+ − 1709
apply(drule mp)
+ − 1710
apply(simp)
+ − 1711
apply(drule mp)
+ − 1712
apply (metis bsimp.elims bsimp.simps(3) good.simps(10) good.simps(11) good.simps(2) good_SEQ)
+ − 1713
apply(simp)
+ − 1714
done
+ − 1715
+ − 1716
+ − 1717
lemma bsimp_idem:
+ − 1718
shows "bsimp (bsimp r) = bsimp r"
+ − 1719
using test good1
+ − 1720
by force
+ − 1721
+ − 1722
+ − 1723
+ − 1724
lemma contains48:
+ − 1725
assumes "\<And>x2aa bs bs1. \<lbrakk>x2aa \<in> set x2a; fuse bs x2aa >> bs @ bs1\<rbrakk> \<Longrightarrow> x2aa >> bs1"
+ − 1726
"AALTs (bs @ x1) x2a >> bs @ bs1"
+ − 1727
shows "AALTs x1 x2a >> bs1"
+ − 1728
using assms
+ − 1729
apply(induct x2a arbitrary: bs x1 bs1)
+ − 1730
apply(auto)
+ − 1731
apply(erule contains.cases)
+ − 1732
apply(auto)
+ − 1733
apply(erule contains.cases)
+ − 1734
apply(auto)
+ − 1735
apply (simp add: contains.intros(4))
+ − 1736
using contains.intros(5) by blast
+ − 1737
+ − 1738
+ − 1739
lemma contains49:
+ − 1740
assumes "fuse bs a >> bs @ bs1"
+ − 1741
shows "a >> bs1"
+ − 1742
using assms
+ − 1743
apply(induct a arbitrary: bs bs1)
+ − 1744
apply(auto)
+ − 1745
using contains.simps apply blast
+ − 1746
apply(erule contains.cases)
+ − 1747
apply(auto)
+ − 1748
apply(rule contains.intros)
+ − 1749
apply(erule contains.cases)
+ − 1750
apply(auto)
+ − 1751
apply(rule contains.intros)
+ − 1752
apply(erule contains.cases)
+ − 1753
apply(auto)
+ − 1754
apply(rule contains.intros)
+ − 1755
apply(auto)[2]
+ − 1756
prefer 2
+ − 1757
apply(erule contains.cases)
+ − 1758
apply(auto)
+ − 1759
apply (simp add: contains.intros(6))
+ − 1760
using contains.intros(7) apply blast
+ − 1761
using contains48 by blast
+ − 1762
+ − 1763
+ − 1764
lemma contains50_IFF2:
+ − 1765
shows "bsimp_AALTs bs [a] >> bs @ bs1 \<longleftrightarrow> fuse bs a >> bs @ bs1"
+ − 1766
by simp
+ − 1767
+ − 1768
lemma contains50_IFF3:
+ − 1769
shows "bsimp_AALTs bs as >> bs @ bs1 \<longleftrightarrow> (\<exists>a \<in> set as. fuse bs a >> bs @ bs1)"
+ − 1770
apply(induct as arbitrary: bs bs1)
+ − 1771
apply(simp)
+ − 1772
apply(auto elim: contains.cases simp add: contains0)
+ − 1773
apply(case_tac as)
+ − 1774
apply(auto)
+ − 1775
apply(case_tac list)
+ − 1776
apply(auto)
+ − 1777
apply(erule contains.cases)
+ − 1778
apply(auto)
+ − 1779
apply (simp add: contains0)
+ − 1780
apply(erule contains.cases)
+ − 1781
apply(auto)
+ − 1782
using contains0 apply auto[1]
+ − 1783
apply(erule contains.cases)
+ − 1784
apply(auto)
+ − 1785
apply(erule contains.cases)
+ − 1786
apply(auto)
+ − 1787
using contains0 apply blast
+ − 1788
apply (metis bsimp_AALTs.simps(2) bsimp_AALTs.simps(3) contains.intros(4) contains49 list.exhaust)
+ − 1789
by (smt bsimp_AALTs.simps(3) contains.intros(4) contains.intros(5) contains49 list.set_cases)
+ − 1790
+ − 1791
lemma contains50_IFF4:
+ − 1792
shows "bsimp_AALTs bs as >> bs @ bs1 \<longleftrightarrow> (\<exists>a \<in> set as. a >> bs1)"
+ − 1793
by (meson contains0 contains49 contains50_IFF3)
+ − 1794
+ − 1795
+ − 1796
lemma contains50:
+ − 1797
assumes "bsimp_AALTs bs rs2 >> bs @ bs1"
+ − 1798
shows "bsimp_AALTs bs (rs1 @ rs2) >> bs @ bs1"
+ − 1799
using assms
+ − 1800
apply(induct rs1 arbitrary: bs rs2 bs1)
+ − 1801
apply(simp)
+ − 1802
apply(auto)
+ − 1803
apply(case_tac rs1)
+ − 1804
apply(simp)
+ − 1805
apply(case_tac rs2)
+ − 1806
apply(simp)
+ − 1807
using contains.simps apply blast
+ − 1808
apply(simp)
+ − 1809
apply(case_tac list)
+ − 1810
apply(simp)
+ − 1811
apply(rule contains.intros)
+ − 1812
back
+ − 1813
apply(rule contains.intros)
+ − 1814
using contains49 apply blast
+ − 1815
apply(simp)
+ − 1816
using contains.intros(5) apply blast
+ − 1817
apply(simp)
+ − 1818
by (metis bsimp_AALTs.elims contains.intros(4) contains.intros(5) contains49 list.distinct(1))
+ − 1819
+ − 1820
lemma contains51:
+ − 1821
assumes "bsimp_AALTs bs [r] >> bs @ bs1"
+ − 1822
shows "bsimp_AALTs bs ([r] @ rs2) >> bs @ bs1"
+ − 1823
using assms
+ − 1824
apply(induct rs2 arbitrary: bs r bs1)
+ − 1825
apply(simp)
+ − 1826
apply(auto)
+ − 1827
using contains.intros(4) contains49 by blast
+ − 1828
+ − 1829
lemma contains51a:
+ − 1830
assumes "bsimp_AALTs bs rs2 >> bs @ bs1"
+ − 1831
shows "bsimp_AALTs bs (rs2 @ [r]) >> bs @ bs1"
+ − 1832
using assms
+ − 1833
apply(induct rs2 arbitrary: bs r bs1)
+ − 1834
apply(simp)
+ − 1835
apply(auto)
+ − 1836
using contains.simps apply blast
+ − 1837
apply(case_tac rs2)
+ − 1838
apply(auto)
+ − 1839
using contains3b contains49 apply blast
+ − 1840
apply(case_tac list)
+ − 1841
apply(auto)
+ − 1842
apply(erule contains.cases)
+ − 1843
apply(auto)
+ − 1844
using contains.intros(4) apply auto[1]
+ − 1845
apply(erule contains.cases)
+ − 1846
apply(auto)
+ − 1847
apply (simp add: contains.intros(4) contains.intros(5))
+ − 1848
apply (simp add: contains.intros(5))
+ − 1849
apply(erule contains.cases)
+ − 1850
apply(auto)
+ − 1851
apply (simp add: contains.intros(4))
+ − 1852
apply(erule contains.cases)
+ − 1853
apply(auto)
+ − 1854
using contains.intros(4) contains.intros(5) apply blast
+ − 1855
using contains.intros(5) by blast
+ − 1856
+ − 1857
lemma contains51b:
+ − 1858
assumes "bsimp_AALTs bs rs >> bs @ bs1"
+ − 1859
shows "bsimp_AALTs bs (rs @ rs2) >> bs @ bs1"
+ − 1860
using assms
+ − 1861
apply(induct rs2 arbitrary: bs rs bs1)
+ − 1862
apply(simp)
+ − 1863
using contains51a by fastforce
+ − 1864
+ − 1865
lemma contains51c:
+ − 1866
assumes "AALTs (bs @ bs2) rs >> bs @ bs1"
+ − 1867
shows "bsimp_AALTs bs (map (fuse bs2) rs) >> bs @ bs1"
+ − 1868
using assms
+ − 1869
apply(induct rs arbitrary: bs bs1 bs2)
+ − 1870
apply(auto)
+ − 1871
apply(erule contains.cases)
+ − 1872
apply(auto)
+ − 1873
apply(erule contains.cases)
+ − 1874
apply(auto)
+ − 1875
using contains0 contains51 apply auto[1]
+ − 1876
by (metis append.left_neutral append_Cons contains50 list.simps(9))
+ − 1877
+ − 1878
+ − 1879
lemma contains51d:
+ − 1880
assumes "fuse bs r >> bs @ bs1"
+ − 1881
shows "bsimp_AALTs bs (flts [r]) >> bs @ bs1"
+ − 1882
using assms
+ − 1883
apply(induct r arbitrary: bs bs1)
+ − 1884
apply(auto)
+ − 1885
by (simp add: contains51c)
+ − 1886
+ − 1887
lemma contains52:
+ − 1888
assumes "\<exists>r \<in> set rs. (fuse bs r) >> bs @ bs1"
+ − 1889
shows "bsimp_AALTs bs (flts rs) >> bs @ bs1"
+ − 1890
using assms
+ − 1891
apply(induct rs arbitrary: bs bs1)
+ − 1892
apply(simp)
+ − 1893
apply(auto)
+ − 1894
defer
+ − 1895
apply (metis contains50 k0)
+ − 1896
apply(subst k0)
+ − 1897
apply(rule contains51b)
+ − 1898
using contains51d by blast
+ − 1899
+ − 1900
lemma contains55:
+ − 1901
assumes "a >> bs"
+ − 1902
shows "bsimp a >> bs"
+ − 1903
using assms
+ − 1904
apply(induct a bs arbitrary:)
+ − 1905
apply(auto intro: contains.intros)
+ − 1906
apply(case_tac "bsimp a1 = AZERO")
+ − 1907
apply(simp)
+ − 1908
using contains.simps apply blast
+ − 1909
apply(case_tac "bsimp a2 = AZERO")
+ − 1910
apply(simp)
+ − 1911
using contains.simps apply blast
+ − 1912
apply(case_tac "\<exists>bs. bsimp a1 = AONE bs")
+ − 1913
apply(auto)[1]
+ − 1914
apply(rotate_tac 1)
+ − 1915
apply(erule contains.cases)
+ − 1916
apply(auto)
+ − 1917
apply (simp add: b1 contains0 fuse_append)
+ − 1918
apply (simp add: bsimp_ASEQ1 contains.intros(3))
+ − 1919
prefer 2
+ − 1920
apply(case_tac rs)
+ − 1921
apply(simp)
+ − 1922
using contains.simps apply blast
+ − 1923
apply (metis contains50 k0)
+ − 1924
(* AALTS case *)
+ − 1925
apply(rule contains52)
+ − 1926
apply(rule_tac x="bsimp r" in bexI)
+ − 1927
apply(auto)
+ − 1928
using contains0 by blast
+ − 1929
+ − 1930
lemma test1:
+ − 1931
shows "AALT [] (ACHAR [Z] c) (ACHAR [S] c) >> [S]"
+ − 1932
by (metis contains.intros(2) contains.intros(4) contains.intros(5) self_append_conv2)
+ − 1933
+ − 1934
lemma test1a:
+ − 1935
shows "bsimp (AALT [] (ACHAR [Z] c) (ACHAR [S] c)) = AALT [] (ACHAR [Z] c) (ACHAR [S] c)"
+ − 1936
apply(simp)
+ − 1937
done
+ − 1938
+ − 1939
lemma q3a:
+ − 1940
assumes "\<exists>r \<in> set rs. bnullable r"
+ − 1941
shows "bmkeps (AALTs bs (map (fuse bs1) rs)) = bmkeps (AALTs (bs@bs1) rs)"
+ − 1942
using assms
+ − 1943
apply(induct rs arbitrary: bs bs1)
+ − 1944
apply(simp)
+ − 1945
apply(simp)
+ − 1946
apply(auto)
+ − 1947
apply (metis append_assoc b2 bnullable_correctness erase_fuse r0)
+ − 1948
apply(case_tac "bnullable a")
+ − 1949
apply (metis append.assoc b2 bnullable_correctness erase_fuse r0)
+ − 1950
apply(case_tac rs)
+ − 1951
apply(simp)
+ − 1952
apply(simp)
+ − 1953
apply(auto)[1]
+ − 1954
apply (metis bnullable_correctness erase_fuse)+
+ − 1955
done
+ − 1956
+ − 1957
+ − 1958
+ − 1959
lemma qq4a:
+ − 1960
assumes "\<exists>x\<in>set list. bnullable x"
+ − 1961
shows "\<exists>x\<in>set (flts list). bnullable x"
+ − 1962
using assms
+ − 1963
apply(induct list rule: flts.induct)
+ − 1964
apply(auto)
+ − 1965
by (metis UnCI bnullable_correctness erase_fuse imageI)
+ − 1966
+ − 1967
+ − 1968
lemma qs3:
+ − 1969
assumes "\<exists>r \<in> set rs. bnullable r"
+ − 1970
shows "bmkeps (AALTs bs rs) = bmkeps (AALTs bs (flts rs))"
+ − 1971
using assms
+ − 1972
apply(induct rs arbitrary: bs taking: size rule: measure_induct)
+ − 1973
apply(case_tac x)
+ − 1974
apply(simp)
+ − 1975
apply(simp)
+ − 1976
apply(case_tac a)
+ − 1977
apply(simp)
+ − 1978
apply (simp add: r1)
+ − 1979
apply(simp)
+ − 1980
apply (simp add: r0)
+ − 1981
apply(simp)
+ − 1982
apply(case_tac "flts list")
+ − 1983
apply(simp)
+ − 1984
apply (metis L_erase_AALTs L_erase_flts L_flat_Prf1 L_flat_Prf2 Prf_elims(1) bnullable_correctness erase.simps(4) mkeps_nullable r2)
+ − 1985
apply(simp)
+ − 1986
apply (simp add: r1)
+ − 1987
prefer 3
+ − 1988
apply(simp)
+ − 1989
apply (simp add: r0)
+ − 1990
prefer 2
+ − 1991
apply(simp)
+ − 1992
apply(case_tac "\<exists>x\<in>set x52. bnullable x")
+ − 1993
apply(case_tac "list")
+ − 1994
apply(simp)
+ − 1995
apply (metis b2 fuse.simps(4) q3a r2)
+ − 1996
apply(erule disjE)
+ − 1997
apply(subst qq1)
+ − 1998
apply(auto)[1]
+ − 1999
apply (metis bnullable_correctness erase_fuse)
+ − 2000
apply(simp)
+ − 2001
apply (metis b2 fuse.simps(4) q3a r2)
+ − 2002
apply(simp)
+ − 2003
apply(auto)[1]
+ − 2004
apply(subst qq1)
+ − 2005
apply (metis bnullable_correctness erase_fuse image_eqI set_map)
+ − 2006
apply (metis b2 fuse.simps(4) q3a r2)
+ − 2007
apply(subst qq1)
+ − 2008
apply (metis bnullable_correctness erase_fuse image_eqI set_map)
+ − 2009
apply (metis b2 fuse.simps(4) q3a r2)
+ − 2010
apply(simp)
+ − 2011
apply(subst qq2)
+ − 2012
apply (metis bnullable_correctness erase_fuse imageE set_map)
+ − 2013
prefer 2
+ − 2014
apply(case_tac "list")
+ − 2015
apply(simp)
+ − 2016
apply(simp)
+ − 2017
apply (simp add: qq4a)
+ − 2018
apply(simp)
+ − 2019
apply(auto)
+ − 2020
apply(case_tac list)
+ − 2021
apply(simp)
+ − 2022
apply(simp)
+ − 2023
apply (simp add: r0)
+ − 2024
apply(case_tac "bnullable (ASEQ x41 x42 x43)")
+ − 2025
apply(case_tac list)
+ − 2026
apply(simp)
+ − 2027
apply(simp)
+ − 2028
apply (simp add: r0)
+ − 2029
apply(simp)
+ − 2030
using qq4a r1 r2 by auto
+ − 2031
+ − 2032
+ − 2033
+ − 2034
lemma k1:
+ − 2035
assumes "\<And>x2aa. \<lbrakk>x2aa \<in> set x2a; bnullable x2aa\<rbrakk> \<Longrightarrow> bmkeps x2aa = bmkeps (bsimp x2aa)"
+ − 2036
"\<exists>x\<in>set x2a. bnullable x"
+ − 2037
shows "bmkeps (AALTs x1 (flts x2a)) = bmkeps (AALTs x1 (flts (map bsimp x2a)))"
+ − 2038
using assms
+ − 2039
apply(induct x2a)
+ − 2040
apply fastforce
+ − 2041
apply(simp)
+ − 2042
apply(subst k0)
+ − 2043
apply(subst (2) k0)
+ − 2044
apply(auto)[1]
+ − 2045
apply (metis b3 k0 list.set_intros(1) qs3 r0)
+ − 2046
by (smt b3 imageI insert_iff k0 list.set(2) qq3 qs3 r0 r1 set_map)
+ − 2047
+ − 2048
+ − 2049
+ − 2050
lemma bmkeps_simp:
+ − 2051
assumes "bnullable r"
+ − 2052
shows "bmkeps r = bmkeps (bsimp r)"
+ − 2053
using assms
+ − 2054
apply(induct r)
+ − 2055
apply(simp)
+ − 2056
apply(simp)
+ − 2057
apply(simp)
+ − 2058
apply(simp)
+ − 2059
prefer 3
+ − 2060
apply(simp)
+ − 2061
apply(case_tac "bsimp r1 = AZERO")
+ − 2062
apply(simp)
+ − 2063
apply(auto)[1]
+ − 2064
apply (metis L_bsimp_erase L_flat_Prf1 L_flat_Prf2 Prf_elims(1) bnullable_correctness erase.simps(1) mkeps_nullable)
+ − 2065
apply(case_tac "bsimp r2 = AZERO")
+ − 2066
apply(simp)
+ − 2067
apply(auto)[1]
+ − 2068
apply (metis L_bsimp_erase L_flat_Prf1 L_flat_Prf2 Prf_elims(1) bnullable_correctness erase.simps(1) mkeps_nullable)
+ − 2069
apply(case_tac "\<exists>bs. bsimp r1 = AONE bs")
+ − 2070
apply(auto)[1]
+ − 2071
apply(subst b1)
+ − 2072
apply(subst b2)
+ − 2073
apply(simp add: b3[symmetric])
+ − 2074
apply(simp)
+ − 2075
apply(subgoal_tac "bsimp_ASEQ x1 (bsimp r1) (bsimp r2) = ASEQ x1 (bsimp r1) (bsimp r2)")
+ − 2076
prefer 2
+ − 2077
apply (smt b3 bnullable.elims(2) bsimp_ASEQ.simps(17) bsimp_ASEQ.simps(19) bsimp_ASEQ.simps(20) bsimp_ASEQ.simps(21) bsimp_ASEQ.simps(22) bsimp_ASEQ.simps(24) bsimp_ASEQ.simps(25) bsimp_ASEQ.simps(26) bsimp_ASEQ.simps(27) bsimp_ASEQ.simps(29) bsimp_ASEQ.simps(30) bsimp_ASEQ.simps(31))
+ − 2078
apply(simp)
+ − 2079
apply(simp)
+ − 2080
thm q3
+ − 2081
apply(subst q3[symmetric])
+ − 2082
apply simp
+ − 2083
using b3 qq4a apply auto[1]
+ − 2084
apply(subst qs3)
+ − 2085
apply simp
+ − 2086
using k1 by blast
+ − 2087
+ − 2088
thm bmkeps_retrieve bmkeps_simp bder_retrieve
+ − 2089
+ − 2090
lemma bmkeps_bder_AALTs:
+ − 2091
assumes "\<exists>r \<in> set rs. bnullable (bder c r)"
+ − 2092
shows "bmkeps (bder c (bsimp_AALTs bs rs)) = bmkeps (bsimp_AALTs bs (map (bder c) rs))"
+ − 2093
using assms
+ − 2094
apply(induct rs)
+ − 2095
apply(simp)
+ − 2096
apply(simp)
+ − 2097
apply(auto)
+ − 2098
apply(case_tac rs)
+ − 2099
apply(simp)
+ − 2100
apply (metis (full_types) Prf_injval bder_retrieve bmkeps_retrieve bnullable_correctness erase_bder erase_fuse mkeps_nullable retrieve_fuse2)
+ − 2101
apply(simp)
+ − 2102
apply(case_tac rs)
+ − 2103
apply(simp_all)
+ − 2104
done
+ − 2105
+ − 2106
lemma bbs0:
+ − 2107
shows "blexer_simp r [] = blexer r []"
+ − 2108
apply(simp add: blexer_def blexer_simp_def)
+ − 2109
done
+ − 2110
+ − 2111
lemma bbs1:
+ − 2112
shows "blexer_simp r [c] = blexer r [c]"
+ − 2113
apply(simp add: blexer_def blexer_simp_def)
+ − 2114
apply(auto)
+ − 2115
defer
+ − 2116
using b3 apply auto[1]
+ − 2117
using b3 apply auto[1]
+ − 2118
apply(subst bmkeps_simp[symmetric])
+ − 2119
apply(simp)
+ − 2120
apply(simp)
+ − 2121
done
+ − 2122
+ − 2123
lemma oo:
+ − 2124
shows "(case (blexer (der c r) s) of None \<Rightarrow> None | Some v \<Rightarrow> Some (injval r c v)) = blexer r (c # s)"
+ − 2125
apply(simp add: blexer_correctness)
+ − 2126
done
+ − 2127
+ − 2128
lemma XXX2_helper:
+ − 2129
assumes "\<forall>y. asize y < Suc (sum_list (map asize rs)) \<longrightarrow> good y \<longrightarrow> bsimp y = y"
+ − 2130
"\<forall>r'\<in>set rs. good r' \<and> nonalt r'"
+ − 2131
shows "flts (map (bsimp \<circ> bder c) (flts (map bsimp rs))) = flts (map (bsimp \<circ> bder c) rs)"
+ − 2132
using assms
+ − 2133
apply(induct rs arbitrary: c)
+ − 2134
apply(simp)
+ − 2135
apply(simp)
+ − 2136
apply(subst k0)
+ − 2137
apply(simp add: flts_append)
+ − 2138
apply(subst (2) k0)
+ − 2139
apply(simp add: flts_append)
+ − 2140
apply(subgoal_tac "flts [a] = [a]")
+ − 2141
prefer 2
+ − 2142
using good.simps(1) k0b apply blast
+ − 2143
apply(simp)
+ − 2144
done
+ − 2145
+ − 2146
lemma bmkeps_good:
+ − 2147
assumes "good a"
+ − 2148
shows "bmkeps (bsimp a) = bmkeps a"
+ − 2149
using assms
+ − 2150
using test2 by auto
+ − 2151
+ − 2152
+ − 2153
lemma xxx_bder:
+ − 2154
assumes "good r"
+ − 2155
shows "L (erase r) \<noteq> {}"
+ − 2156
using assms
+ − 2157
apply(induct r rule: good.induct)
+ − 2158
apply(auto simp add: Sequ_def)
+ − 2159
done
+ − 2160
+ − 2161
lemma xxx_bder2:
+ − 2162
assumes "L (erase (bsimp r)) = {}"
+ − 2163
shows "bsimp r = AZERO"
+ − 2164
using assms xxx_bder test2 good1
+ − 2165
by blast
+ − 2166
+ − 2167
lemma XXX2aa:
+ − 2168
assumes "good a"
+ − 2169
shows "bsimp (bder c (bsimp a)) = bsimp (bder c a)"
+ − 2170
using assms
+ − 2171
by (simp add: test2)
+ − 2172
+ − 2173
lemma XXX2aa_ders:
+ − 2174
assumes "good a"
+ − 2175
shows "bsimp (bders (bsimp a) s) = bsimp (bders a s)"
+ − 2176
using assms
+ − 2177
by (simp add: test2)
+ − 2178
+ − 2179
lemma XXX4a:
+ − 2180
shows "good (bders_simp (bsimp r) s) \<or> bders_simp (bsimp r) s = AZERO"
+ − 2181
apply(induct s arbitrary: r rule: rev_induct)
+ − 2182
apply(simp)
+ − 2183
apply (simp add: good1)
+ − 2184
apply(simp add: bders_simp_append)
+ − 2185
apply (simp add: good1)
+ − 2186
done
+ − 2187
+ − 2188
lemma XXX4a_good:
+ − 2189
assumes "good a"
+ − 2190
shows "good (bders_simp a s) \<or> bders_simp a s = AZERO"
+ − 2191
using assms
+ − 2192
apply(induct s arbitrary: a rule: rev_induct)
+ − 2193
apply(simp)
+ − 2194
apply(simp add: bders_simp_append)
+ − 2195
apply (simp add: good1)
+ − 2196
done
+ − 2197
+ − 2198
lemma XXX4a_good_cons:
+ − 2199
assumes "s \<noteq> []"
+ − 2200
shows "good (bders_simp a s) \<or> bders_simp a s = AZERO"
+ − 2201
using assms
+ − 2202
apply(case_tac s)
+ − 2203
apply(auto)
+ − 2204
using XXX4a by blast
+ − 2205
+ − 2206
lemma XXX4b:
+ − 2207
assumes "good a" "L (erase (bders_simp a s)) \<noteq> {}"
+ − 2208
shows "good (bders_simp a s)"
+ − 2209
using assms
+ − 2210
apply(induct s arbitrary: a)
+ − 2211
apply(simp)
+ − 2212
apply(simp)
+ − 2213
apply(subgoal_tac "L (erase (bder a aa)) = {} \<or> L (erase (bder a aa)) \<noteq> {}")
+ − 2214
prefer 2
+ − 2215
apply(auto)[1]
+ − 2216
apply(erule disjE)
+ − 2217
apply(subgoal_tac "bsimp (bder a aa) = AZERO")
+ − 2218
prefer 2
+ − 2219
using L_bsimp_erase xxx_bder2 apply auto[1]
+ − 2220
apply(simp)
+ − 2221
apply (metis L.simps(1) XXX4a erase.simps(1))
+ − 2222
apply(drule_tac x="bsimp (bder a aa)" in meta_spec)
+ − 2223
apply(drule meta_mp)
+ − 2224
apply simp
+ − 2225
apply(rule good1a)
+ − 2226
apply(auto)
+ − 2227
done
+ − 2228
+ − 2229
lemma bders_AZERO:
+ − 2230
shows "bders AZERO s = AZERO"
+ − 2231
and "bders_simp AZERO s = AZERO"
+ − 2232
apply (induct s)
+ − 2233
apply(auto)
+ − 2234
done
+ − 2235
+ − 2236
lemma LA:
+ − 2237
assumes "\<Turnstile> v : ders s (erase r)"
+ − 2238
shows "retrieve (bders r s) v = retrieve r (flex (erase r) id s v)"
+ − 2239
using assms
+ − 2240
apply(induct s arbitrary: r v rule: rev_induct)
+ − 2241
apply(simp)
+ − 2242
apply(simp add: bders_append ders_append)
+ − 2243
apply(subst bder_retrieve)
+ − 2244
apply(simp)
+ − 2245
apply(drule Prf_injval)
+ − 2246
by (simp add: flex_append)
+ − 2247
+ − 2248
+ − 2249
lemma LB:
+ − 2250
assumes "s \<in> (erase r) \<rightarrow> v"
+ − 2251
shows "retrieve r v = retrieve r (flex (erase r) id s (mkeps (ders s (erase r))))"
+ − 2252
using assms
+ − 2253
apply(induct s arbitrary: r v rule: rev_induct)
+ − 2254
apply(simp)
+ − 2255
apply(subgoal_tac "v = mkeps (erase r)")
+ − 2256
prefer 2
+ − 2257
apply (simp add: Posix1(1) Posix_determ Posix_mkeps nullable_correctness)
+ − 2258
apply(simp)
+ − 2259
apply(simp add: flex_append ders_append)
+ − 2260
by (metis Posix_determ Posix_flex Posix_injval Posix_mkeps ders_snoc lexer_correctness(2) lexer_flex)
+ − 2261
+ − 2262
lemma LB_sym:
+ − 2263
assumes "s \<in> (erase r) \<rightarrow> v"
+ − 2264
shows "retrieve r v = retrieve r (flex (erase r) id s (mkeps (erase (bders r s))))"
+ − 2265
using assms
+ − 2266
by (simp add: LB)
+ − 2267
+ − 2268
+ − 2269
lemma LC:
+ − 2270
assumes "s \<in> (erase r) \<rightarrow> v"
+ − 2271
shows "retrieve r v = retrieve (bders r s) (mkeps (erase (bders r s)))"
+ − 2272
apply(simp)
+ − 2273
by (metis LA LB Posix1(1) assms lexer_correct_None lexer_flex mkeps_nullable)
+ − 2274
+ − 2275
+ − 2276
lemma L0:
+ − 2277
assumes "bnullable a"
+ − 2278
shows "retrieve (bsimp a) (mkeps (erase (bsimp a))) = retrieve a (mkeps (erase a))"
+ − 2279
using assms b3 bmkeps_retrieve bmkeps_simp bnullable_correctness
+ − 2280
by (metis b3 bmkeps_retrieve bmkeps_simp bnullable_correctness)
+ − 2281
+ − 2282
thm bmkeps_retrieve
+ − 2283
+ − 2284
lemma L0a:
+ − 2285
assumes "s \<in> L(erase a)"
+ − 2286
shows "retrieve (bsimp (bders a s)) (mkeps (erase (bsimp (bders a s)))) =
+ − 2287
retrieve (bders a s) (mkeps (erase (bders a s)))"
+ − 2288
using assms
+ − 2289
by (metis L0 bnullable_correctness erase_bders lexer_correct_None lexer_flex)
+ − 2290
+ − 2291
lemma L0aa:
+ − 2292
assumes "s \<in> L (erase a)"
+ − 2293
shows "[] \<in> erase (bsimp (bders a s)) \<rightarrow> mkeps (erase (bsimp (bders a s)))"
+ − 2294
using assms
+ − 2295
by (metis Posix_mkeps b3 bnullable_correctness erase_bders lexer_correct_None lexer_flex)
+ − 2296
+ − 2297
lemma L0aaa:
+ − 2298
assumes "[c] \<in> L (erase a)"
+ − 2299
shows "[c] \<in> (erase a) \<rightarrow> flex (erase a) id [c] (mkeps (erase (bder c a)))"
+ − 2300
using assms
+ − 2301
by (metis bders.simps(1) bders.simps(2) erase_bders lexer_correct_None lexer_correct_Some lexer_flex option.inject)
+ − 2302
+ − 2303
lemma L0aaaa:
+ − 2304
assumes "[c] \<in> L (erase a)"
+ − 2305
shows "[c] \<in> (erase a) \<rightarrow> flex (erase a) id [c] (mkeps (erase (bders a [c])))"
+ − 2306
using assms
+ − 2307
using L0aaa by auto
+ − 2308
+ − 2309
+ − 2310
lemma L02:
+ − 2311
assumes "bnullable (bder c a)"
+ − 2312
shows "retrieve (bsimp a) (flex (erase (bsimp a)) id [c] (mkeps (erase (bder c (bsimp a))))) =
+ − 2313
retrieve (bder c (bsimp a)) (mkeps (erase (bder c (bsimp a))))"
+ − 2314
using assms
+ − 2315
apply(simp)
+ − 2316
using bder_retrieve L0 bmkeps_simp bmkeps_retrieve L0 LA LB
+ − 2317
apply(subst bder_retrieve[symmetric])
+ − 2318
apply (metis L_bsimp_erase bnullable_correctness der_correctness erase_bder mkeps_nullable nullable_correctness)
+ − 2319
apply(simp)
+ − 2320
done
+ − 2321
+ − 2322
lemma L02_bders:
+ − 2323
assumes "bnullable (bders a s)"
+ − 2324
shows "retrieve (bsimp a) (flex (erase (bsimp a)) id s (mkeps (erase (bders (bsimp a) s)))) =
+ − 2325
retrieve (bders (bsimp a) s) (mkeps (erase (bders (bsimp a) s)))"
+ − 2326
using assms
+ − 2327
by (metis LA L_bsimp_erase bnullable_correctness ders_correctness erase_bders mkeps_nullable nullable_correctness)
+ − 2328
+ − 2329
+ − 2330
+ − 2331
+ − 2332
lemma L03:
+ − 2333
assumes "bnullable (bder c a)"
+ − 2334
shows "retrieve (bder c (bsimp a)) (mkeps (erase (bder c (bsimp a)))) =
+ − 2335
bmkeps (bsimp (bder c (bsimp a)))"
+ − 2336
using assms
+ − 2337
by (metis L0 L_bsimp_erase bmkeps_retrieve bnullable_correctness der_correctness erase_bder nullable_correctness)
+ − 2338
+ − 2339
lemma L04:
+ − 2340
assumes "bnullable (bder c a)"
+ − 2341
shows "retrieve (bder c (bsimp a)) (mkeps (erase (bder c (bsimp a)))) =
+ − 2342
retrieve (bsimp (bder c (bsimp a))) (mkeps (erase (bsimp (bder c (bsimp a)))))"
+ − 2343
using assms
+ − 2344
by (metis L0 L_bsimp_erase bnullable_correctness der_correctness erase_bder nullable_correctness)
+ − 2345
+ − 2346
lemma L05:
+ − 2347
assumes "bnullable (bder c a)"
+ − 2348
shows "retrieve (bder c (bsimp a)) (mkeps (erase (bder c (bsimp a)))) =
+ − 2349
retrieve (bsimp (bder c (bsimp a))) (mkeps (erase (bsimp (bder c (bsimp a)))))"
+ − 2350
using assms
+ − 2351
using L04 by auto
+ − 2352
+ − 2353
lemma L06:
+ − 2354
assumes "bnullable (bder c a)"
+ − 2355
shows "bmkeps (bder c (bsimp a)) = bmkeps (bsimp (bder c (bsimp a)))"
+ − 2356
using assms
+ − 2357
by (metis L03 L_bsimp_erase bmkeps_retrieve bnullable_correctness der_correctness erase_bder nullable_correctness)
+ − 2358
+ − 2359
lemma L07:
+ − 2360
assumes "s \<in> L (erase r)"
+ − 2361
shows "retrieve r (flex (erase r) id s (mkeps (ders s (erase r))))
+ − 2362
= retrieve (bders r s) (mkeps (erase (bders r s)))"
+ − 2363
using assms
+ − 2364
using LB LC lexer_correct_Some by auto
+ − 2365
+ − 2366
lemma L06_2:
+ − 2367
assumes "bnullable (bders a [c,d])"
+ − 2368
shows "bmkeps (bders (bsimp a) [c,d]) = bmkeps (bsimp (bders (bsimp a) [c,d]))"
+ − 2369
using assms
+ − 2370
apply(simp)
+ − 2371
by (metis L_bsimp_erase bmkeps_simp bnullable_correctness der_correctness erase_bder nullable_correctness)
+ − 2372
+ − 2373
lemma L06_bders:
+ − 2374
assumes "bnullable (bders a s)"
+ − 2375
shows "bmkeps (bders (bsimp a) s) = bmkeps (bsimp (bders (bsimp a) s))"
+ − 2376
using assms
+ − 2377
by (metis L_bsimp_erase bmkeps_simp bnullable_correctness ders_correctness erase_bders nullable_correctness)
+ − 2378
+ − 2379
lemma LLLL:
+ − 2380
shows "L (erase a) = L (erase (bsimp a))"
+ − 2381
and "L (erase a) = {flat v | v. \<Turnstile> v: (erase a)}"
+ − 2382
and "L (erase a) = {flat v | v. \<Turnstile> v: (erase (bsimp a))}"
+ − 2383
using L_bsimp_erase apply(blast)
+ − 2384
apply (simp add: L_flat_Prf)
+ − 2385
using L_bsimp_erase L_flat_Prf apply(auto)[1]
+ − 2386
done
+ − 2387
+ − 2388
+ − 2389
+ − 2390
lemma L07XX:
+ − 2391
assumes "s \<in> L (erase a)"
+ − 2392
shows "s \<in> erase a \<rightarrow> flex (erase a) id s (mkeps (ders s (erase a)))"
+ − 2393
using assms
+ − 2394
by (meson lexer_correct_None lexer_correctness(1) lexer_flex)
+ − 2395
+ − 2396
lemma LX0:
+ − 2397
assumes "s \<in> L r"
+ − 2398
shows "decode (bmkeps (bders (intern r) s)) r = Some(flex r id s (mkeps (ders s r)))"
+ − 2399
by (metis assms blexer_correctness blexer_def lexer_correct_None lexer_flex)
+ − 2400
+ − 2401
lemma L1:
+ − 2402
assumes "s \<in> r \<rightarrow> v"
+ − 2403
shows "decode (bmkeps (bders (intern r) s)) r = Some v"
+ − 2404
using assms
+ − 2405
by (metis blexer_correctness blexer_def lexer_correctness(1) option.distinct(1))
+ − 2406
+ − 2407
lemma L2:
+ − 2408
assumes "s \<in> (der c r) \<rightarrow> v"
+ − 2409
shows "decode (bmkeps (bders (intern r) (c # s))) r = Some (injval r c v)"
+ − 2410
using assms
+ − 2411
apply(subst bmkeps_retrieve)
+ − 2412
using Posix1(1) lexer_correct_None lexer_flex apply fastforce
+ − 2413
using MAIN_decode
+ − 2414
apply(subst MAIN_decode[symmetric])
+ − 2415
apply(simp)
+ − 2416
apply (meson Posix1(1) lexer_correct_None lexer_flex mkeps_nullable)
+ − 2417
apply(simp)
+ − 2418
apply(subgoal_tac "v = flex (der c r) id s (mkeps (ders s (der c r)))")
+ − 2419
prefer 2
+ − 2420
apply (metis Posix_determ lexer_correctness(1) lexer_flex option.distinct(1))
+ − 2421
apply(simp)
+ − 2422
apply(subgoal_tac "injval r c (flex (der c r) id s (mkeps (ders s (der c r)))) =
+ − 2423
(flex (der c r) ((\<lambda>v. injval r c v) o id) s (mkeps (ders s (der c r))))")
+ − 2424
apply(simp)
+ − 2425
using flex_fun_apply by blast
+ − 2426
+ − 2427
lemma L3:
+ − 2428
assumes "s2 \<in> (ders s1 r) \<rightarrow> v"
+ − 2429
shows "decode (bmkeps (bders (intern r) (s1 @ s2))) r = Some (flex r id s1 v)"
+ − 2430
using assms
+ − 2431
apply(induct s1 arbitrary: r s2 v rule: rev_induct)
+ − 2432
apply(simp)
+ − 2433
using L1 apply blast
+ − 2434
apply(simp add: ders_append)
+ − 2435
apply(drule_tac x="r" in meta_spec)
+ − 2436
apply(drule_tac x="x # s2" in meta_spec)
+ − 2437
apply(drule_tac x="injval (ders xs r) x v" in meta_spec)
+ − 2438
apply(drule meta_mp)
+ − 2439
defer
+ − 2440
apply(simp)
+ − 2441
apply(simp add: flex_append)
+ − 2442
by (simp add: Posix_injval)
+ − 2443
+ − 2444
+ − 2445
+ − 2446
lemma bders_snoc:
+ − 2447
"bder c (bders a s) = bders a (s @ [c])"
+ − 2448
apply(simp add: bders_append)
+ − 2449
done
+ − 2450
+ − 2451
+ − 2452
lemma QQ1:
+ − 2453
shows "bsimp (bders (bsimp a) []) = bders_simp (bsimp a) []"
+ − 2454
apply(simp)
+ − 2455
apply(simp add: bsimp_idem)
+ − 2456
done
+ − 2457
+ − 2458
lemma QQ2:
+ − 2459
shows "bsimp (bders (bsimp a) [c]) = bders_simp (bsimp a) [c]"
+ − 2460
apply(simp)
+ − 2461
done
+ − 2462
+ − 2463
lemma XXX2a_long:
+ − 2464
assumes "good a"
+ − 2465
shows "bsimp (bder c (bsimp a)) = bsimp (bder c a)"
+ − 2466
using assms
+ − 2467
apply(induct a arbitrary: c taking: asize rule: measure_induct)
+ − 2468
apply(case_tac x)
+ − 2469
apply(simp)
+ − 2470
apply(simp)
+ − 2471
apply(simp)
+ − 2472
prefer 3
+ − 2473
apply(simp)
+ − 2474
apply(simp)
+ − 2475
apply(auto)[1]
+ − 2476
apply(case_tac "x42 = AZERO")
+ − 2477
apply(simp)
+ − 2478
apply(case_tac "x43 = AZERO")
+ − 2479
apply(simp)
+ − 2480
using test2 apply force
+ − 2481
apply(case_tac "\<exists>bs. x42 = AONE bs")
+ − 2482
apply(clarify)
+ − 2483
apply(simp)
+ − 2484
apply(subst bsimp_ASEQ1)
+ − 2485
apply(simp)
+ − 2486
using b3 apply force
+ − 2487
using bsimp_ASEQ0 test2 apply force
+ − 2488
thm good_SEQ test2
+ − 2489
apply (simp add: good_SEQ test2)
+ − 2490
apply (simp add: good_SEQ test2)
+ − 2491
apply(case_tac "x42 = AZERO")
+ − 2492
apply(simp)
+ − 2493
apply(case_tac "x43 = AZERO")
+ − 2494
apply(simp)
+ − 2495
apply (simp add: bsimp_ASEQ0)
+ − 2496
apply(case_tac "\<exists>bs. x42 = AONE bs")
+ − 2497
apply(clarify)
+ − 2498
apply(simp)
+ − 2499
apply(subst bsimp_ASEQ1)
+ − 2500
apply(simp)
+ − 2501
using bsimp_ASEQ0 test2 apply force
+ − 2502
apply (simp add: good_SEQ test2)
+ − 2503
apply (simp add: good_SEQ test2)
+ − 2504
apply (simp add: good_SEQ test2)
+ − 2505
(* AALTs case *)
+ − 2506
apply(simp)
+ − 2507
using test2 by fastforce
+ − 2508
+ − 2509
+ − 2510
lemma bder_bsimp_AALTs:
+ − 2511
shows "bder c (bsimp_AALTs bs rs) = bsimp_AALTs bs (map (bder c) rs)"
+ − 2512
apply(induct bs rs rule: bsimp_AALTs.induct)
+ − 2513
apply(simp)
+ − 2514
apply(simp)
+ − 2515
apply (simp add: bder_fuse)
+ − 2516
apply(simp)
+ − 2517
done
+ − 2518
+ − 2519
lemma bders_bsimp_AALTs:
+ − 2520
shows "bders (bsimp_AALTs bs rs) s = bsimp_AALTs bs (map (\<lambda>a. bders a s) rs)"
+ − 2521
apply(induct s arbitrary: bs rs rule: rev_induct)
+ − 2522
apply(simp)
+ − 2523
apply(simp add: bders_append)
+ − 2524
apply(simp add: bder_bsimp_AALTs)
+ − 2525
apply(simp add: comp_def)
+ − 2526
done
+ − 2527
+ − 2528
lemma flts_nothing:
+ − 2529
assumes "\<forall>r \<in> set rs. r \<noteq> AZERO" "\<forall>r \<in> set rs. nonalt r"
+ − 2530
shows "flts rs = rs"
+ − 2531
using assms
+ − 2532
apply(induct rs rule: flts.induct)
+ − 2533
apply(auto)
+ − 2534
done
+ − 2535
+ − 2536
lemma flts_flts:
+ − 2537
assumes "\<forall>r \<in> set rs. good r"
+ − 2538
shows "flts (flts rs) = flts rs"
+ − 2539
using assms
+ − 2540
apply(induct rs taking: "\<lambda>rs. sum_list (map asize rs)" rule: measure_induct)
+ − 2541
apply(case_tac x)
+ − 2542
apply(simp)
+ − 2543
apply(simp)
+ − 2544
apply(case_tac a)
+ − 2545
apply(simp_all add: bder_fuse flts_append)
+ − 2546
apply(subgoal_tac "\<forall>r \<in> set x52. r \<noteq> AZERO")
+ − 2547
prefer 2
+ − 2548
apply (metis Nil_is_append_conv bsimp_AALTs.elims good.simps(1) good.simps(5) good0 list.distinct(1) n0 nn1b split_list_last test2)
+ − 2549
apply(subgoal_tac "\<forall>r \<in> set x52. nonalt r")
+ − 2550
prefer 2
+ − 2551
apply (metis n0 nn1b test2)
+ − 2552
by (metis flts_fuse flts_nothing)
+ − 2553
+ − 2554
+ − 2555
lemma iii:
+ − 2556
assumes "bsimp_AALTs bs rs \<noteq> AZERO"
+ − 2557
shows "rs \<noteq> []"
+ − 2558
using assms
+ − 2559
apply(induct bs rs rule: bsimp_AALTs.induct)
+ − 2560
apply(auto)
+ − 2561
done
+ − 2562
+ − 2563
lemma CT1_SEQ:
+ − 2564
shows "bsimp (ASEQ bs a1 a2) = bsimp (ASEQ bs (bsimp a1) (bsimp a2))"
+ − 2565
apply(simp add: bsimp_idem)
+ − 2566
done
+ − 2567
+ − 2568
lemma CT1:
+ − 2569
shows "bsimp (AALTs bs as) = bsimp (AALTs bs (map bsimp as))"
+ − 2570
apply(induct as arbitrary: bs)
+ − 2571
apply(simp)
+ − 2572
apply(simp)
+ − 2573
by (simp add: bsimp_idem comp_def)
+ − 2574
+ − 2575
lemma CT1a:
+ − 2576
shows "bsimp (AALT bs a1 a2) = bsimp(AALT bs (bsimp a1) (bsimp a2))"
+ − 2577
by (metis CT1 list.simps(8) list.simps(9))
+ − 2578
+ − 2579
lemma WWW2:
+ − 2580
shows "bsimp (bsimp_AALTs bs1 (flts (map bsimp as1))) =
+ − 2581
bsimp_AALTs bs1 (flts (map bsimp as1))"
+ − 2582
by (metis bsimp.simps(2) bsimp_idem)
+ − 2583
+ − 2584
lemma CT1b:
+ − 2585
shows "bsimp (bsimp_AALTs bs as) = bsimp (bsimp_AALTs bs (map bsimp as))"
+ − 2586
apply(induct bs as rule: bsimp_AALTs.induct)
+ − 2587
apply(auto simp add: bsimp_idem)
+ − 2588
apply (simp add: bsimp_fuse bsimp_idem)
+ − 2589
by (metis bsimp_idem comp_apply)
+ − 2590
+ − 2591
+ − 2592
+ − 2593
+ − 2594
(* CT *)
+ − 2595
+ − 2596
lemma CTa:
+ − 2597
assumes "\<forall>r \<in> set as. nonalt r \<and> r \<noteq> AZERO"
+ − 2598
shows "flts as = as"
+ − 2599
using assms
+ − 2600
apply(induct as)
+ − 2601
apply(simp)
+ − 2602
apply(case_tac as)
+ − 2603
apply(simp)
+ − 2604
apply (simp add: k0b)
+ − 2605
using flts_nothing by auto
+ − 2606
+ − 2607
lemma CT0:
+ − 2608
assumes "\<forall>r \<in> set as1. nonalt r \<and> r \<noteq> AZERO"
+ − 2609
shows "flts [bsimp_AALTs bs1 as1] = flts (map (fuse bs1) as1)"
+ − 2610
using assms CTa
+ − 2611
apply(induct as1 arbitrary: bs1)
+ − 2612
apply(simp)
+ − 2613
apply(simp)
+ − 2614
apply(case_tac as1)
+ − 2615
apply(simp)
+ − 2616
apply(simp)
+ − 2617
proof -
+ − 2618
fix a :: arexp and as1a :: "arexp list" and bs1a :: "bit list" and aa :: arexp and list :: "arexp list"
+ − 2619
assume a1: "nonalt a \<and> a \<noteq> AZERO \<and> nonalt aa \<and> aa \<noteq> AZERO \<and> (\<forall>r\<in>set list. nonalt r \<and> r \<noteq> AZERO)"
+ − 2620
assume a2: "\<And>as. \<forall>r\<in>set as. nonalt r \<and> r \<noteq> AZERO \<Longrightarrow> flts as = as"
+ − 2621
assume a3: "as1a = aa # list"
+ − 2622
have "flts [a] = [a]"
+ − 2623
using a1 k0b by blast
+ − 2624
then show "fuse bs1a a # fuse bs1a aa # map (fuse bs1a) list = flts (fuse bs1a a # fuse bs1a aa # map (fuse bs1a) list)"
+ − 2625
using a3 a2 a1 by (metis (no_types) append.left_neutral append_Cons flts_fuse k00 k0b list.simps(9))
+ − 2626
qed
+ − 2627
+ − 2628
+ − 2629
lemma CT01:
+ − 2630
assumes "\<forall>r \<in> set as1. nonalt r \<and> r \<noteq> AZERO" "\<forall>r \<in> set as2. nonalt r \<and> r \<noteq> AZERO"
+ − 2631
shows "flts [bsimp_AALTs bs1 as1, bsimp_AALTs bs2 as2] = flts ((map (fuse bs1) as1) @ (map (fuse bs2) as2))"
+ − 2632
using assms CT0
+ − 2633
by (metis k0 k00)
+ − 2634
+ − 2635
+ − 2636
+ − 2637
lemma CT_exp:
+ − 2638
assumes "\<forall>a \<in> set as. bsimp (bder c (bsimp a)) = bsimp (bder c a)"
+ − 2639
shows "map bsimp (map (bder c) as) = map bsimp (map (bder c) (map bsimp as))"
+ − 2640
using assms
+ − 2641
apply(induct as)
+ − 2642
apply(auto)
+ − 2643
done
+ − 2644
+ − 2645
lemma asize_set:
+ − 2646
assumes "a \<in> set as"
+ − 2647
shows "asize a < Suc (sum_list (map asize as))"
+ − 2648
using assms
+ − 2649
apply(induct as arbitrary: a)
+ − 2650
apply(auto)
+ − 2651
using le_add2 le_less_trans not_less_eq by blast
+ − 2652
+ − 2653
lemma L_erase_bder_simp:
+ − 2654
shows "L (erase (bsimp (bder a r))) = L (der a (erase (bsimp r)))"
+ − 2655
using L_bsimp_erase der_correctness by auto
+ − 2656
+ − 2657
lemma PPP0:
+ − 2658
assumes "s \<in> r \<rightarrow> v"
+ − 2659
shows "(bders (intern r) s) >> code v"
+ − 2660
using assms
+ − 2661
by (smt L07 L1 LX0 Posix1(1) Posix_Prf contains6 erase_bders erase_intern lexer_correct_None lexer_flex mkeps_nullable option.inject retrieve_code)
+ − 2662
+ − 2663
thm L07 L1 LX0 Posix1(1) Posix_Prf contains6 erase_bders erase_intern lexer_correct_None lexer_flex mkeps_nullable option.inject retrieve_code
+ − 2664
+ − 2665
+ − 2666
lemma PPP0_isar:
+ − 2667
assumes "s \<in> r \<rightarrow> v"
+ − 2668
shows "(bders (intern r) s) >> code v"
+ − 2669
proof -
+ − 2670
from assms have a1: "\<Turnstile> v : r" using Posix_Prf by simp
+ − 2671
+ − 2672
from assms have "s \<in> L r" using Posix1(1) by auto
+ − 2673
then have "[] \<in> L (ders s r)" by (simp add: ders_correctness Ders_def)
+ − 2674
then have a2: "\<Turnstile> mkeps (ders s r) : ders s r"
+ − 2675
by (simp add: mkeps_nullable nullable_correctness)
+ − 2676
+ − 2677
have "retrieve (bders (intern r) s) (mkeps (ders s r)) =
+ − 2678
retrieve (intern r) (flex r id s (mkeps (ders s r)))" using a2 LA LB bder_retrieve by simp
+ − 2679
also have "... = retrieve (intern r) v"
+ − 2680
using LB assms by auto
+ − 2681
also have "... = code v" using a1 by (simp add: retrieve_code)
+ − 2682
finally have "retrieve (bders (intern r) s) (mkeps (ders s r)) = code v" by simp
+ − 2683
moreover
+ − 2684
have "\<Turnstile> mkeps (ders s r) : erase (bders (intern r) s)" using a2 by simp
+ − 2685
then have "bders (intern r) s >> retrieve (bders (intern r) s) (mkeps (ders s r))"
+ − 2686
by (rule contains6)
+ − 2687
ultimately
+ − 2688
show "(bders (intern r) s) >> code v" by simp
+ − 2689
qed
+ − 2690
+ − 2691
lemma PPP0b:
+ − 2692
assumes "s \<in> r \<rightarrow> v"
+ − 2693
shows "(intern r) >> code v"
+ − 2694
using assms
+ − 2695
using Posix_Prf contains2 by auto
+ − 2696
+ − 2697
lemma PPP0_eq:
+ − 2698
assumes "s \<in> r \<rightarrow> v"
+ − 2699
shows "(intern r >> code v) = (bders (intern r) s >> code v)"
+ − 2700
using assms
+ − 2701
using PPP0_isar PPP0b by blast
+ − 2702
+ − 2703
lemma f_cont1:
+ − 2704
assumes "fuse bs1 a >> bs"
+ − 2705
shows "\<exists>bs2. bs = bs1 @ bs2"
+ − 2706
using assms
+ − 2707
apply(induct a arbitrary: bs1 bs)
+ − 2708
apply(auto elim: contains.cases)
+ − 2709
done
+ − 2710
+ − 2711
+ − 2712
lemma f_cont2:
+ − 2713
assumes "bsimp_AALTs bs1 as >> bs"
+ − 2714
shows "\<exists>bs2. bs = bs1 @ bs2"
+ − 2715
using assms
+ − 2716
apply(induct bs1 as arbitrary: bs rule: bsimp_AALTs.induct)
+ − 2717
apply(auto elim: contains.cases f_cont1)
+ − 2718
done
+ − 2719
+ − 2720
lemma contains_SEQ1:
+ − 2721
assumes "bsimp_ASEQ bs r1 r2 >> bsX"
+ − 2722
shows "\<exists>bs1 bs2. r1 >> bs1 \<and> r2 >> bs2 \<and> bsX = bs @ bs1 @ bs2"
+ − 2723
using assms
+ − 2724
apply(auto)
+ − 2725
apply(case_tac "r1 = AZERO")
+ − 2726
apply(auto)
+ − 2727
using contains.simps apply blast
+ − 2728
apply(case_tac "r2 = AZERO")
+ − 2729
apply(auto)
+ − 2730
apply(simp add: bsimp_ASEQ0)
+ − 2731
using contains.simps apply blast
+ − 2732
apply(case_tac "\<exists>bsX. r1 = AONE bsX")
+ − 2733
apply(auto)
+ − 2734
apply(simp add: bsimp_ASEQ2)
+ − 2735
apply (metis append_assoc contains.intros(1) contains49 f_cont1)
+ − 2736
apply(simp add: bsimp_ASEQ1)
+ − 2737
apply(erule contains.cases)
+ − 2738
apply(auto)
+ − 2739
done
+ − 2740
+ − 2741
lemma contains59:
+ − 2742
assumes "AALTs bs rs >> bs2"
+ − 2743
shows "\<exists>r \<in> set rs. (fuse bs r) >> bs2"
+ − 2744
using assms
+ − 2745
apply(induct rs arbitrary: bs bs2)
+ − 2746
apply(auto)
+ − 2747
apply(erule contains.cases)
+ − 2748
apply(auto)
+ − 2749
apply(erule contains.cases)
+ − 2750
apply(auto)
+ − 2751
using contains0 by blast
+ − 2752
+ − 2753
lemma contains60:
+ − 2754
assumes "\<exists>r \<in> set rs. fuse bs r >> bs2"
+ − 2755
shows "AALTs bs rs >> bs2"
+ − 2756
using assms
+ − 2757
apply(induct rs arbitrary: bs bs2)
+ − 2758
apply(auto)
+ − 2759
apply (metis contains3b contains49 f_cont1)
+ − 2760
using contains.intros(5) f_cont1 by blast
+ − 2761
+ − 2762
+ − 2763
+ − 2764
lemma contains61:
+ − 2765
assumes "bsimp_AALTs bs rs >> bs2"
+ − 2766
shows "\<exists>r \<in> set rs. (fuse bs r) >> bs2"
+ − 2767
using assms
+ − 2768
apply(induct arbitrary: bs2 rule: bsimp_AALTs.induct)
+ − 2769
apply(auto)
+ − 2770
using contains.simps apply blast
+ − 2771
using contains59 by fastforce
+ − 2772
+ − 2773
lemma contains61b:
+ − 2774
assumes "bsimp_AALTs bs rs >> bs2"
+ − 2775
shows "\<exists>r \<in> set (flts rs). (fuse bs r) >> bs2"
+ − 2776
using assms
+ − 2777
apply(induct bs rs arbitrary: bs2 rule: bsimp_AALTs.induct)
+ − 2778
apply(auto)
+ − 2779
using contains.simps apply blast
+ − 2780
using contains51d contains61 f_cont1 apply blast
+ − 2781
by (metis bsimp_AALTs.simps(3) contains52 contains61 f_cont2)
+ − 2782
+ − 2783
+ − 2784
+ − 2785
lemma contains61a:
+ − 2786
assumes "\<exists>r \<in> set rs. (fuse bs r) >> bs2"
+ − 2787
shows "bsimp_AALTs bs rs >> bs2"
+ − 2788
using assms
+ − 2789
apply(induct rs arbitrary: bs2 bs)
+ − 2790
apply(auto)
+ − 2791
apply (metis bsimp_AALTs.elims contains60 list.distinct(1) list.inject list.set_intros(1))
+ − 2792
by (metis append_Cons append_Nil contains50 f_cont2)
+ − 2793
+ − 2794
lemma contains62:
+ − 2795
assumes "bsimp_AALTs bs (rs1 @ rs2) >> bs2"
+ − 2796
shows "bsimp_AALTs bs rs1 >> bs2 \<or> bsimp_AALTs bs rs2 >> bs2"
+ − 2797
using assms
+ − 2798
apply -
+ − 2799
apply(drule contains61)
+ − 2800
apply(auto)
+ − 2801
apply(case_tac rs1)
+ − 2802
apply(auto)
+ − 2803
apply(case_tac list)
+ − 2804
apply(auto)
+ − 2805
apply (simp add: contains60)
+ − 2806
apply(case_tac list)
+ − 2807
apply(auto)
+ − 2808
apply (simp add: contains60)
+ − 2809
apply (meson contains60 list.set_intros(2))
+ − 2810
apply(case_tac rs2)
+ − 2811
apply(auto)
+ − 2812
apply(case_tac list)
+ − 2813
apply(auto)
+ − 2814
apply (simp add: contains60)
+ − 2815
apply(case_tac list)
+ − 2816
apply(auto)
+ − 2817
apply (simp add: contains60)
+ − 2818
apply (meson contains60 list.set_intros(2))
+ − 2819
done
+ − 2820
+ − 2821
lemma contains63:
+ − 2822
assumes "AALTs bs (map (fuse bs1) rs) >> bs3"
+ − 2823
shows "AALTs (bs @ bs1) rs >> bs3"
+ − 2824
using assms
+ − 2825
apply(induct rs arbitrary: bs bs1 bs3)
+ − 2826
apply(auto elim: contains.cases)
+ − 2827
apply(erule contains.cases)
+ − 2828
apply(auto)
+ − 2829
apply (simp add: contains0 contains60 fuse_append)
+ − 2830
by (metis contains.intros(5) contains59 f_cont1)
+ − 2831
+ − 2832
lemma contains64:
+ − 2833
assumes "bsimp_AALTs bs (flts rs1 @ flts rs2) >> bs2" "\<forall>r \<in> set rs2. \<not> fuse bs r >> bs2"
+ − 2834
shows "bsimp_AALTs bs (flts rs1) >> bs2"
+ − 2835
using assms
+ − 2836
apply(induct rs2 arbitrary: rs1 bs bs2)
+ − 2837
apply(auto)
+ − 2838
apply(drule_tac x="rs1" in meta_spec)
+ − 2839
apply(drule_tac x="bs" in meta_spec)
+ − 2840
apply(drule_tac x="bs2" in meta_spec)
+ − 2841
apply(drule meta_mp)
+ − 2842
apply(drule contains61)
+ − 2843
apply(auto)
+ − 2844
using contains51b contains61a f_cont1 apply blast
+ − 2845
apply(subst (asm) k0)
+ − 2846
apply(auto)
+ − 2847
prefer 2
+ − 2848
using contains50 contains61a f_cont1 apply blast
+ − 2849
apply(case_tac a)
+ − 2850
apply(auto)
+ − 2851
by (metis contains60 fuse_append)
+ − 2852
+ − 2853
+ − 2854
+ − 2855
lemma contains65:
+ − 2856
assumes "bsimp_AALTs bs (flts rs) >> bs2"
+ − 2857
shows "\<exists>r \<in> set rs. (fuse bs r) >> bs2"
+ − 2858
using assms
+ − 2859
apply(induct rs arbitrary: bs bs2 taking: "\<lambda>rs. sum_list (map asize rs)" rule: measure_induct)
+ − 2860
apply(case_tac x)
+ − 2861
apply(auto elim: contains.cases)
+ − 2862
apply(case_tac list)
+ − 2863
apply(auto elim: contains.cases)
+ − 2864
apply(case_tac a)
+ − 2865
apply(auto elim: contains.cases)
+ − 2866
apply(drule contains61)
+ − 2867
apply(auto)
+ − 2868
apply (metis contains60 fuse_append)
+ − 2869
apply(case_tac lista)
+ − 2870
apply(auto elim: contains.cases)
+ − 2871
apply(subst (asm) k0)
+ − 2872
apply(drule contains62)
+ − 2873
apply(auto)
+ − 2874
apply(case_tac a)
+ − 2875
apply(auto elim: contains.cases)
+ − 2876
apply(case_tac x52)
+ − 2877
apply(auto elim: contains.cases)
+ − 2878
apply(case_tac list)
+ − 2879
apply(auto elim: contains.cases)
+ − 2880
apply (simp add: contains60 fuse_append)
+ − 2881
apply(erule contains.cases)
+ − 2882
apply(auto)
+ − 2883
apply (metis append.left_neutral contains0 contains60 fuse.simps(4) in_set_conv_decomp)
+ − 2884
apply(erule contains.cases)
+ − 2885
apply(auto)
+ − 2886
apply (metis contains0 contains60 fuse.simps(4) list.set_intros(1) list.set_intros(2))
+ − 2887
apply (simp add: contains.intros(5) contains63)
+ − 2888
apply(case_tac aa)
+ − 2889
apply(auto)
+ − 2890
apply (meson contains60 contains61 contains63)
+ − 2891
apply(subst (asm) k0)
+ − 2892
apply(drule contains64)
+ − 2893
apply(auto)[1]
+ − 2894
by (metis append_Nil2 bsimp_AALTs.simps(2) contains50 contains61a contains64 f_cont2 flts.simps(1))
+ − 2895
+ − 2896
+ − 2897
lemma contains55a:
+ − 2898
assumes "bsimp r >> bs"
+ − 2899
shows "r >> bs"
+ − 2900
using assms
+ − 2901
apply(induct r arbitrary: bs)
+ − 2902
apply(auto)
+ − 2903
apply(frule contains_SEQ1)
+ − 2904
apply(auto)
+ − 2905
apply (simp add: contains.intros(3))
+ − 2906
apply(frule f_cont2)
+ − 2907
apply(auto)
+ − 2908
apply(drule contains65)
+ − 2909
apply(auto)
+ − 2910
using contains0 contains49 contains60 by blast
+ − 2911
+ − 2912
+ − 2913
lemma PPP1_eq:
+ − 2914
shows "bsimp r >> bs \<longleftrightarrow> r >> bs"
+ − 2915
using contains55 contains55a by blast
+ − 2916
+ − 2917
+ − 2918
definition "SET a \<equiv> {bs . a >> bs}"
+ − 2919
+ − 2920
lemma "SET(bsimp a) \<subseteq> SET(a)"
+ − 2921
unfolding SET_def
+ − 2922
apply(auto simp add: PPP1_eq)
+ − 2923
done
+ − 2924
+ − 2925
lemma retrieve_code_bder:
+ − 2926
assumes "\<Turnstile> v : der c r"
+ − 2927
shows "code (injval r c v) = retrieve (bder c (intern r)) v"
+ − 2928
using assms
+ − 2929
by (simp add: Prf_injval bder_retrieve retrieve_code)
+ − 2930
+ − 2931
lemma Etrans:
+ − 2932
assumes "a >> s" "s = t"
+ − 2933
shows "a >> t"
+ − 2934
using assms by simp
+ − 2935
+ − 2936
+ − 2937
+ − 2938
lemma retrieve_code_bders:
+ − 2939
assumes "\<Turnstile> v : ders s r"
+ − 2940
shows "code (flex r id s v) = retrieve (bders (intern r) s) v"
+ − 2941
using assms
+ − 2942
apply(induct s arbitrary: v r rule: rev_induct)
+ − 2943
apply(auto simp add: ders_append flex_append bders_append)
+ − 2944
apply (simp add: retrieve_code)
+ − 2945
apply(frule Prf_injval)
+ − 2946
apply(drule_tac meta_spec)+
+ − 2947
apply(drule meta_mp)
+ − 2948
apply(assumption)
+ − 2949
apply(simp)
+ − 2950
apply(subst bder_retrieve)
+ − 2951
apply(simp)
+ − 2952
apply(simp)
+ − 2953
done
+ − 2954
+ − 2955
lemma contains70:
+ − 2956
assumes "s \<in> L(r)"
+ − 2957
shows "bders (intern r) s >> code (flex r id s (mkeps (ders s r)))"
+ − 2958
apply(subst PPP0_eq[symmetric])
+ − 2959
apply (meson assms lexer_correct_None lexer_correctness(1) lexer_flex)
+ − 2960
by (metis L07XX PPP0b assms erase_intern)
+ − 2961
+ − 2962
+ − 2963
+ − 2964
lemma PPP:
+ − 2965
assumes "\<Turnstile> v : r"
+ − 2966
shows "intern r >> (retrieve (intern r) v)"
+ − 2967
using assms
+ − 2968
using contains5 by blast
+ − 2969
+ − 2970
+ − 2971
+ − 2972
+ − 2973
+ − 2974
+ − 2975
+ − 2976
+ − 2977
definition FC where
+ − 2978
"FC a s v = retrieve a (flex (erase a) id s v)"
+ − 2979
+ − 2980
definition FE where
+ − 2981
"FE a s = retrieve a (flex (erase a) id s (mkeps (ders s (erase a))))"
+ − 2982
+ − 2983
definition PV where
+ − 2984
"PV r s v = flex r id s v"
+ − 2985
+ − 2986
definition PX where
+ − 2987
"PX r s = PV r s (mkeps (ders s r))"
+ − 2988
+ − 2989
+ − 2990
lemma FE_PX:
+ − 2991
shows "FE r s = retrieve r (PX (erase r) s)"
+ − 2992
unfolding FE_def PX_def PV_def by(simp)
+ − 2993
+ − 2994
lemma FE_PX_code:
+ − 2995
assumes "s \<in> L r"
+ − 2996
shows "FE (intern r) s = code (PX r s)"
+ − 2997
unfolding FE_def PX_def PV_def
+ − 2998
using assms
+ − 2999
by (metis L07XX Posix_Prf erase_intern retrieve_code)
+ − 3000
+ − 3001
+ − 3002
lemma PV_id[simp]:
+ − 3003
shows "PV r [] v = v"
+ − 3004
by (simp add: PV_def)
+ − 3005
+ − 3006
lemma PX_id[simp]:
+ − 3007
shows "PX r [] = mkeps r"
+ − 3008
by (simp add: PX_def)
+ − 3009
+ − 3010
lemma PV_cons:
+ − 3011
shows "PV r (c # s) v = injval r c (PV (der c r) s v)"
+ − 3012
apply(simp add: PV_def flex_fun_apply)
+ − 3013
done
+ − 3014
+ − 3015
lemma PX_cons:
+ − 3016
shows "PX r (c # s) = injval r c (PX (der c r) s)"
+ − 3017
apply(simp add: PX_def PV_cons)
+ − 3018
done
+ − 3019
+ − 3020
lemma PV_append:
+ − 3021
shows "PV r (s1 @ s2) v = PV r s1 (PV (ders s1 r) s2 v)"
+ − 3022
apply(simp add: PV_def flex_append)
+ − 3023
by (simp add: flex_fun_apply2)
+ − 3024
+ − 3025
lemma PX_append:
+ − 3026
shows "PX r (s1 @ s2) = PV r s1 (PX (ders s1 r) s2)"
+ − 3027
by (simp add: PV_append PX_def ders_append)
+ − 3028
+ − 3029
lemma code_PV0:
+ − 3030
shows "PV r (c # s) v = injval r c (PV (der c r) s v)"
+ − 3031
unfolding PX_def PV_def
+ − 3032
apply(simp)
+ − 3033
by (simp add: flex_injval)
+ − 3034
+ − 3035
lemma code_PX0:
+ − 3036
shows "PX r (c # s) = injval r c (PX (der c r) s)"
+ − 3037
unfolding PX_def
+ − 3038
apply(simp add: code_PV0)
+ − 3039
done
+ − 3040
+ − 3041
lemma Prf_PV:
+ − 3042
assumes "\<Turnstile> v : ders s r"
+ − 3043
shows "\<Turnstile> PV r s v : r"
+ − 3044
using assms unfolding PX_def PV_def
+ − 3045
apply(induct s arbitrary: v r)
+ − 3046
apply(simp)
+ − 3047
apply(simp)
+ − 3048
by (simp add: Prf_injval flex_injval)
+ − 3049
+ − 3050
+ − 3051
lemma Prf_PX:
+ − 3052
assumes "s \<in> L r"
+ − 3053
shows "\<Turnstile> PX r s : r"
+ − 3054
using assms unfolding PX_def PV_def
+ − 3055
using L1 LX0 Posix_Prf lexer_correct_Some by fastforce
+ − 3056
+ − 3057
lemma PV1:
+ − 3058
assumes "\<Turnstile> v : ders s r"
+ − 3059
shows "(intern r) >> code (PV r s v)"
+ − 3060
using assms
+ − 3061
by (simp add: Prf_PV contains2)
+ − 3062
+ − 3063
lemma PX1:
+ − 3064
assumes "s \<in> L r"
+ − 3065
shows "(intern r) >> code (PX r s)"
+ − 3066
using assms
+ − 3067
by (simp add: Prf_PX contains2)
+ − 3068
+ − 3069
lemma PX2:
+ − 3070
assumes "s \<in> L (der c r)"
+ − 3071
shows "bder c (intern r) >> code (injval r c (PX (der c r) s))"
+ − 3072
using assms
+ − 3073
by (simp add: Prf_PX contains6 retrieve_code_bder)
+ − 3074
+ − 3075
lemma PX2a:
+ − 3076
assumes "c # s \<in> L r"
+ − 3077
shows "bder c (intern r) >> code (injval r c (PX (der c r) s))"
+ − 3078
using assms
+ − 3079
using PX2 lexer_correct_None by force
+ − 3080
+ − 3081
lemma PX2b:
+ − 3082
assumes "c # s \<in> L r"
+ − 3083
shows "bder c (intern r) >> code (PX r (c # s))"
+ − 3084
using assms unfolding PX_def PV_def
+ − 3085
by (metis Der_def L07XX PV_def PX2a PX_def Posix_determ Posix_injval der_correctness erase_intern mem_Collect_eq)
+ − 3086
+ − 3087
lemma PV3:
+ − 3088
assumes "\<Turnstile> v : ders s r"
+ − 3089
shows "bders (intern r) s >> code (PV r s v)"
+ − 3090
using assms
+ − 3091
using PX_def PV_def contains70
+ − 3092
by (simp add: contains6 retrieve_code_bders)
+ − 3093
+ − 3094
lemma PX3:
+ − 3095
assumes "s \<in> L r"
+ − 3096
shows "bders (intern r) s >> code (PX r s)"
+ − 3097
using assms
+ − 3098
using PX_def PV_def contains70 by auto
+ − 3099
+ − 3100
+ − 3101
lemma PV_bders_iff:
+ − 3102
assumes "\<Turnstile> v : ders s r"
+ − 3103
shows "bders (intern r) s >> code (PV r s v) \<longleftrightarrow> (intern r) >> code (PV r s v)"
+ − 3104
by (simp add: PV1 PV3 assms)
+ − 3105
+ − 3106
lemma PX_bders_iff:
+ − 3107
assumes "s \<in> L r"
+ − 3108
shows "bders (intern r) s >> code (PX r s) \<longleftrightarrow> (intern r) >> code (PX r s)"
+ − 3109
by (simp add: PX1 PX3 assms)
+ − 3110
+ − 3111
lemma PX4:
+ − 3112
assumes "(s1 @ s2) \<in> L r"
+ − 3113
shows "bders (intern r) (s1 @ s2) >> code (PX r (s1 @ s2))"
+ − 3114
using assms
+ − 3115
by (simp add: PX3)
+ − 3116
+ − 3117
lemma PX_bders_iff2:
+ − 3118
assumes "(s1 @ s2) \<in> L r"
+ − 3119
shows "bders (intern r) (s1 @ s2) >> code (PX r (s1 @ s2)) \<longleftrightarrow>
+ − 3120
(intern r) >> code (PX r (s1 @ s2))"
+ − 3121
by (simp add: PX1 PX3 assms)
+ − 3122
+ − 3123
lemma PV_bders_iff3:
+ − 3124
assumes "\<Turnstile> v : ders (s1 @ s2) r"
+ − 3125
shows "bders (intern r) (s1 @ s2) >> code (PV r (s1 @ s2) v) \<longleftrightarrow>
+ − 3126
bders (intern r) s1 >> code (PV r (s1 @ s2) v)"
+ − 3127
by (metis PV3 PV_append Prf_PV assms ders_append)
+ − 3128
+ − 3129
+ − 3130
+ − 3131
lemma PX_bders_iff3:
+ − 3132
assumes "(s1 @ s2) \<in> L r"
+ − 3133
shows "bders (intern r) (s1 @ s2) >> code (PX r (s1 @ s2)) \<longleftrightarrow>
+ − 3134
bders (intern r) s1 >> code (PX r (s1 @ s2))"
+ − 3135
by (metis Ders_def L07XX PV_append PV_def PX4 PX_def Posix_Prf assms contains6 ders_append ders_correctness erase_bders erase_intern mem_Collect_eq retrieve_code_bders)
+ − 3136
+ − 3137
lemma PV_bder_iff:
+ − 3138
assumes "\<Turnstile> v : ders (s1 @ [c]) r"
+ − 3139
shows "bder c (bders (intern r) s1) >> code (PV r (s1 @ [c]) v) \<longleftrightarrow>
+ − 3140
bders (intern r) s1 >> code (PV r (s1 @ [c]) v)"
+ − 3141
by (simp add: PV_bders_iff3 assms bders_snoc)
+ − 3142
+ − 3143
lemma PV_bder_IFF:
+ − 3144
assumes "\<Turnstile> v : ders (s1 @ c # s2) r"
+ − 3145
shows "bder c (bders (intern r) s1) >> code (PV r (s1 @ c # s2) v) \<longleftrightarrow>
+ − 3146
bders (intern r) s1 >> code (PV r (s1 @ c # s2) v)"
+ − 3147
by (metis LA PV3 PV_def Prf_PV assms bders_append code_PV0 contains7 ders.simps(2) erase_bders erase_intern retrieve_code_bders)
+ − 3148
+ − 3149
+ − 3150
lemma PX_bder_iff:
+ − 3151
assumes "(s1 @ [c]) \<in> L r"
+ − 3152
shows "bder c (bders (intern r) s1) >> code (PX r (s1 @ [c])) \<longleftrightarrow>
+ − 3153
bders (intern r) s1 >> code (PX r (s1 @ [c]))"
+ − 3154
by (simp add: PX_bders_iff3 assms bders_snoc)
+ − 3155
+ − 3156
lemma PV_bder_iff2:
+ − 3157
assumes "\<Turnstile> v : ders (c # s1) r"
+ − 3158
shows "bders (bder c (intern r)) s1 >> code (PV r (c # s1) v) \<longleftrightarrow>
+ − 3159
bder c (intern r) >> code (PV r (c # s1) v)"
+ − 3160
by (metis PV3 Prf_PV assms bders.simps(2) code_PV0 contains7 ders.simps(2) erase_intern retrieve_code)
+ − 3161
+ − 3162
+ − 3163
lemma PX_bder_iff2:
+ − 3164
assumes "(c # s1) \<in> L r"
+ − 3165
shows "bders (bder c (intern r)) s1 >> code (PX r (c # s1)) \<longleftrightarrow>
+ − 3166
bder c (intern r) >> code (PX r (c # s1))"
+ − 3167
using PX2b PX3 assms by force
+ − 3168
+ − 3169
+ − 3170
lemma FC_id:
+ − 3171
shows "FC r [] v = retrieve r v"
+ − 3172
by (simp add: FC_def)
+ − 3173
+ − 3174
lemma FC_char:
+ − 3175
shows "FC r [c] v = retrieve r (injval (erase r) c v)"
+ − 3176
by (simp add: FC_def)
+ − 3177
+ − 3178
lemma FC_char2:
+ − 3179
assumes "\<Turnstile> v : der c (erase r)"
+ − 3180
shows "FC r [c] v = FC (bder c r) [] v"
+ − 3181
using assms
+ − 3182
by (simp add: FC_char FC_id bder_retrieve)
+ − 3183
+ − 3184
+ − 3185
lemma FC_bders_iff:
+ − 3186
assumes "\<Turnstile> v : ders s (erase r)"
+ − 3187
shows "bders r s >> FC r s v \<longleftrightarrow> r >> FC r s v"
+ − 3188
unfolding FC_def
+ − 3189
by (simp add: assms contains8_iff)
+ − 3190
+ − 3191
+ − 3192
lemma FC_bder_iff:
+ − 3193
assumes "\<Turnstile> v : der c (erase r)"
+ − 3194
shows "bder c r >> FC r [c] v \<longleftrightarrow> r >> FC r [c] v"
+ − 3195
apply(subst FC_bders_iff[symmetric])
+ − 3196
apply(simp add: assms)
+ − 3197
apply(simp)
+ − 3198
done
+ − 3199
+ − 3200
lemma FC_bders_iff2:
+ − 3201
assumes "\<Turnstile> v : ders (c # s) (erase r)"
+ − 3202
shows "bders r (c # s) >> FC r (c # s) v \<longleftrightarrow> bders (bder c r) s >> FC (bder c r) s v"
+ − 3203
apply(subst FC_bders_iff)
+ − 3204
using assms apply simp
+ − 3205
by (metis FC_def assms contains7b contains8_iff ders.simps(2) erase_bder)
+ − 3206
+ − 3207
+ − 3208
lemma FC_bnullable0:
+ − 3209
assumes "bnullable r"
+ − 3210
shows "FC r [] (mkeps (erase r)) = FC (bsimp r) [] (mkeps (erase (bsimp r)))"
+ − 3211
unfolding FC_def
+ − 3212
by (simp add: L0 assms)
+ − 3213
+ − 3214
+ − 3215
lemma FC_nullable2:
+ − 3216
assumes "bnullable (bders a s)"
+ − 3217
shows "FC (bsimp a) s (mkeps (erase (bders (bsimp a) s))) =
+ − 3218
FC (bders (bsimp a) s) [] (mkeps (erase (bders (bsimp a) s)))"
+ − 3219
unfolding FC_def
+ − 3220
using L02_bders assms by auto
+ − 3221
+ − 3222
lemma FC_nullable3:
+ − 3223
assumes "bnullable (bders a s)"
+ − 3224
shows "FC a s (mkeps (erase (bders a s))) =
+ − 3225
FC (bders a s) [] (mkeps (erase (bders a s)))"
+ − 3226
unfolding FC_def
+ − 3227
using LA assms bnullable_correctness mkeps_nullable by fastforce
+ − 3228
+ − 3229
+ − 3230
lemma FE_contains0:
+ − 3231
assumes "bnullable r"
+ − 3232
shows "r >> FE r []"
+ − 3233
by (simp add: FE_def assms bnullable_correctness contains6 mkeps_nullable)
+ − 3234
+ − 3235
lemma FE_contains1:
+ − 3236
assumes "bnullable (bders r s)"
+ − 3237
shows "r >> FE r s"
+ − 3238
by (metis FE_def Prf_flex assms bnullable_correctness contains6 erase_bders mkeps_nullable)
+ − 3239
+ − 3240
lemma FE_bnullable0:
+ − 3241
assumes "bnullable r"
+ − 3242
shows "FE r [] = FE (bsimp r) []"
+ − 3243
unfolding FE_def
+ − 3244
by (simp add: L0 assms)
+ − 3245
+ − 3246
+ − 3247
lemma FE_nullable1:
+ − 3248
assumes "bnullable (bders r s)"
+ − 3249
shows "FE r s = FE (bders r s) []"
+ − 3250
unfolding FE_def
+ − 3251
using LA assms bnullable_correctness mkeps_nullable by fastforce
+ − 3252
+ − 3253
lemma FE_contains2:
+ − 3254
assumes "bnullable (bders r s)"
+ − 3255
shows "r >> FE (bders r s) []"
+ − 3256
by (metis FE_contains1 FE_nullable1 assms)
+ − 3257
+ − 3258
lemma FE_contains3:
+ − 3259
assumes "bnullable (bder c r)"
+ − 3260
shows "r >> FE (bsimp (bder c r)) []"
+ − 3261
by (metis FE_def L0 assms bder_retrieve bders.simps(1) bnullable_correctness contains7a erase_bder erase_bders flex.simps(1) id_apply mkeps_nullable)
+ − 3262
+ − 3263
lemma FE_contains4:
+ − 3264
assumes "bnullable (bders r s)"
+ − 3265
shows "r >> FE (bsimp (bders r s)) []"
+ − 3266
using FE_bnullable0 FE_contains2 assms by auto
+ − 3267
+ − 3268
lemma FC4:
+ − 3269
assumes "\<Turnstile> v : ders s (erase a)"
+ − 3270
shows "FC a s v = FC (bders a s) [] v"
+ − 3271
unfolding FC_def by (simp add: LA assms)
+ − 3272
+ − 3273
lemma FC5:
+ − 3274
assumes "nullable (erase a)"
+ − 3275
shows "FC a [] (mkeps (erase a)) = FC (bsimp a) [] (mkeps (erase (bsimp a)))"
+ − 3276
unfolding FC_def
+ − 3277
using L0 assms bnullable_correctness by auto
+ − 3278
+ − 3279
+ − 3280
lemma in1:
+ − 3281
assumes "AALTs bsX rsX \<in> set rs"
+ − 3282
shows "\<forall>r \<in> set rsX. fuse bsX r \<in> set (flts rs)"
+ − 3283
using assms
+ − 3284
apply(induct rs arbitrary: bsX rsX)
+ − 3285
apply(auto)
+ − 3286
by (metis append_assoc in_set_conv_decomp k0)
+ − 3287
+ − 3288
lemma in2a:
+ − 3289
assumes "nonnested (bsimp r)" "\<not>nonalt(bsimp r)"
+ − 3290
shows "(\<exists>bsX rsX. r = AALTs bsX rsX) \<or> (\<exists>bsX rX1 rX2. r = ASEQ bsX rX1 rX2 \<and> bnullable rX1)"
+ − 3291
using assms
+ − 3292
apply(induct r)
+ − 3293
apply(auto)
+ − 3294
by (metis arexp.distinct(25) b3 bnullable.simps(2) bsimp_ASEQ.simps(1) bsimp_ASEQ0 bsimp_ASEQ1 nonalt.elims(3) nonalt.simps(2))
+ − 3295
+ − 3296
+ − 3297
lemma [simp]:
+ − 3298
shows "size (fuse bs r) = size r"
+ − 3299
by (induct r) (auto)
+ − 3300
+ − 3301
fun AALTs_subs where
+ − 3302
"AALTs_subs (AZERO) = {}"
+ − 3303
| "AALTs_subs (AONE bs) = {AONE bs}"
+ − 3304
| "AALTs_subs (ACHAR bs c) = {ACHAR bs c}"
+ − 3305
| "AALTs_subs (ASEQ bs r1 r2) = {ASEQ bs r1 r2}"
+ − 3306
| "AALTs_subs (ASTAR bs r) = {ASTAR bs r}"
+ − 3307
| "AALTs_subs (AALTs bs []) = {}"
+ − 3308
| "AALTs_subs (AALTs bs (r#rs)) = AALTs_subs (fuse bs r) \<union> AALTs_subs (AALTs bs rs)"
+ − 3309
+ − 3310
lemma nonalt_10:
+ − 3311
assumes "nonalt r" "r \<noteq> AZERO"
+ − 3312
shows "r \<in> AALTs_subs r"
+ − 3313
using assms
+ − 3314
apply(induct r)
+ − 3315
apply(auto)
+ − 3316
done
+ − 3317
+ − 3318
lemma flt_fuse:
+ − 3319
shows "flts (map (fuse bs) rs) = map (fuse bs) (flts rs)"
+ − 3320
apply(induct rs arbitrary: bs rule: flts.induct)
+ − 3321
apply(auto)
+ − 3322
by (simp add: fuse_append)
+ − 3323
+ − 3324
lemma AALTs_subs_fuse:
+ − 3325
shows "AALTs_subs (fuse bs r) = (fuse bs) ` (AALTs_subs r)"
+ − 3326
apply(induct r arbitrary: bs rule: AALTs_subs.induct)
+ − 3327
apply(auto)
+ − 3328
apply (simp add: fuse_append)
+ − 3329
apply blast
+ − 3330
by (simp add: fuse_append)
+ − 3331
+ − 3332
lemma AALTs_subs_fuse2:
+ − 3333
shows "AALTs_subs (AALTs bs rs) = AALTs_subs (AALTs [] (map (fuse bs) rs))"
+ − 3334
apply(induct rs arbitrary: bs)
+ − 3335
apply(auto)
+ − 3336
apply (auto simp add: fuse_empty)
+ − 3337
done
+ − 3338
+ − 3339
lemma fuse_map:
+ − 3340
shows "map (fuse (bs1 @ bs2)) rs = map (fuse bs1) (map (fuse bs2) rs)"
+ − 3341
apply(induct rs)
+ − 3342
apply(auto)
+ − 3343
using fuse_append by blast
+ − 3344
+ − 3345
+ − 3346
+ − 3347
lemma contains59_2:
+ − 3348
assumes "AALTs bs rs >> bs2"
+ − 3349
shows "\<exists>r\<in>AALTs_subs (AALTs bs rs). r >> bs2"
+ − 3350
using assms
+ − 3351
apply(induct rs arbitrary: bs bs2 taking: "\<lambda>rs. sum_list (map asize rs)" rule: measure_induct)
+ − 3352
apply(case_tac x)
+ − 3353
apply(auto)
+ − 3354
using contains59 apply force
+ − 3355
apply(erule contains.cases)
+ − 3356
apply(auto)
+ − 3357
apply(case_tac "r = AZERO")
+ − 3358
apply(simp)
+ − 3359
apply (metis bsimp_AALTs.simps(1) contains61 empty_iff empty_set)
+ − 3360
apply(case_tac "nonalt r")
+ − 3361
apply (metis UnCI bsimp_AALTs.simps(1) contains0 contains61 empty_iff empty_set nn11a nonalt_10)
+ − 3362
apply(subgoal_tac "\<exists>bsX rsX. r = AALTs bsX rsX")
+ − 3363
prefer 2
+ − 3364
using bbbbs1 apply blast
+ − 3365
apply(auto)
+ − 3366
apply (metis UnCI contains0 fuse.simps(4) less_add_Suc1)
+ − 3367
apply(drule_tac x="rs" in spec)
+ − 3368
apply(drule mp)
+ − 3369
apply(simp add: asize0)
+ − 3370
apply(drule_tac x="bsa" in spec)
+ − 3371
apply(drule_tac x="bsa @ bs1" in spec)
+ − 3372
apply(auto)
+ − 3373
done
+ − 3374
+ − 3375
lemma TEMPLATE_contains61a:
+ − 3376
assumes "\<exists>r \<in> set rs. (fuse bs r) >> bs2"
+ − 3377
shows "bsimp_AALTs bs rs >> bs2"
+ − 3378
using assms
+ − 3379
apply(induct rs arbitrary: bs2 bs)
+ − 3380
apply(auto)
+ − 3381
apply (metis bsimp_AALTs.elims contains60 list.distinct(1) list.inject list.set_intros(1))
+ − 3382
by (metis append_Cons append_Nil contains50 f_cont2)
+ − 3383
+ − 3384
+ − 3385
+ − 3386
+ − 3387
lemma H1:
+ − 3388
assumes "r >> bs2" "r \<in> AALTs_subs a"
+ − 3389
shows "a >> bs2"
+ − 3390
using assms
+ − 3391
apply(induct a arbitrary: r bs2 rule: AALTs_subs.induct)
+ − 3392
apply(auto)
+ − 3393
apply (simp add: contains60)
+ − 3394
by (simp add: contains59 contains60)
+ − 3395
+ − 3396
lemma H3:
+ − 3397
assumes "a >> bs"
+ − 3398
shows "\<exists>r \<in> AALTs_subs a. r >> bs"
+ − 3399
using assms
+ − 3400
apply(induct a bs)
+ − 3401
apply(auto intro: contains.intros)
+ − 3402
using contains.intros(4) contains59_2 by fastforce
+ − 3403
+ − 3404
lemma H4:
+ − 3405
shows "AALTs_subs (AALTs bs rs1) \<subseteq> AALTs_subs (AALTs bs (rs1 @ rs2))"
+ − 3406
apply(induct rs1)
+ − 3407
apply(auto)
+ − 3408
done
+ − 3409
+ − 3410
lemma H5:
+ − 3411
shows "AALTs_subs (AALTs bs rs2) \<subseteq> AALTs_subs (AALTs bs (rs1 @ rs2))"
+ − 3412
apply(induct rs1)
+ − 3413
apply(auto)
+ − 3414
done
+ − 3415
+ − 3416
lemma H7:
+ − 3417
shows "AALTs_subs (AALTs bs (rs1 @ rs2)) = AALTs_subs (AALTs bs rs1) \<union> AALTs_subs (AALTs bs rs2)"
+ − 3418
apply(induct rs1)
+ − 3419
apply(auto)
+ − 3420
done
+ − 3421
+ − 3422
lemma H10:
+ − 3423
shows "AALTs_subs (AALTs bs rs) = (\<Union>r \<in> set rs. AALTs_subs (fuse bs r))"
+ − 3424
apply(induct rs arbitrary: bs)
+ − 3425
apply(auto)
+ − 3426
done
+ − 3427
+ − 3428
lemma H6:
+ − 3429
shows "AALTs_subs (AALTs bs (flts rs)) = AALTs_subs (AALTs bs rs)"
+ − 3430
apply(induct rs arbitrary: bs rule: flts.induct)
+ − 3431
apply(auto)
+ − 3432
apply (metis AALTs_subs_fuse2 H7 Un_iff fuse_map)
+ − 3433
apply (metis AALTs_subs_fuse2 H7 UnCI fuse_map)
+ − 3434
by (simp add: H7)
+ − 3435
+ − 3436
+ − 3437
+ − 3438
lemma H2:
+ − 3439
assumes "r >> bs2" "r \<in> AALTs_subs (AALTs bs rs)"
+ − 3440
shows "r \<in> AALTs_subs (AALTs bs (flts rs))"
+ − 3441
using assms
+ − 3442
apply(induct rs arbitrary: r bs bs2 rule: flts.induct)
+ − 3443
apply(auto)
+ − 3444
apply (metis AALTs_subs_fuse2 H4 fuse_map in_mono)
+ − 3445
using H7 by blast
+ − 3446
+ − 3447
lemma HH1:
+ − 3448
assumes "r \<in> AALTs_subs (fuse bs a)" "r >> bs2"
+ − 3449
shows "\<exists>bs3. bs2 = bs @ bs3"
+ − 3450
using assms
+ − 3451
using H1 f_cont1 by blast
+ − 3452
+ − 3453
lemma fuse_inj:
+ − 3454
assumes "fuse bs a = fuse bs b"
+ − 3455
shows "a = b"
+ − 3456
using assms
+ − 3457
apply(induct a arbitrary: bs b)
+ − 3458
apply(auto)
+ − 3459
apply(case_tac b)
+ − 3460
apply(auto)
+ − 3461
apply(case_tac b)
+ − 3462
apply(auto)
+ − 3463
apply(case_tac b)
+ − 3464
apply(auto)
+ − 3465
apply(case_tac b)
+ − 3466
apply(auto)
+ − 3467
apply(case_tac b)
+ − 3468
apply(auto)
+ − 3469
apply(case_tac b)
+ − 3470
apply(auto)
+ − 3471
done
+ − 3472
+ − 3473
lemma HH11:
+ − 3474
assumes "r \<in> AALTs_subs (fuse bs1 a)"
+ − 3475
shows "fuse bs r \<in> AALTs_subs (fuse (bs @ bs1) a)"
+ − 3476
using assms
+ − 3477
apply(induct a arbitrary: r bs bs1)
+ − 3478
apply(auto)
+ − 3479
apply(subst (asm) H10)
+ − 3480
apply(auto)
+ − 3481
apply(drule_tac x="x" in meta_spec)
+ − 3482
apply(simp)
+ − 3483
apply(drule_tac x="r" in meta_spec)
+ − 3484
apply(drule_tac x="bs" in meta_spec)
+ − 3485
apply(drule_tac x="bs1 @ x1" in meta_spec)
+ − 3486
apply(simp)
+ − 3487
apply(subst H10)
+ − 3488
apply(auto)
+ − 3489
done
+ − 3490
+ − 3491
lemma HH12:
+ − 3492
assumes "r \<in> AALTs_subs a"
+ − 3493
shows "fuse bs r \<in> AALTs_subs (fuse bs a)"
+ − 3494
using AALTs_subs_fuse assms by blast
+ − 3495
+ − 3496
lemma HH13:
+ − 3497
assumes "r \<in> (\<Union>r \<in> set rs. AALTs_subs r)"
+ − 3498
shows "fuse bs r \<in> AALTs_subs (AALTs bs rs)"
+ − 3499
using assms
+ − 3500
using H10 HH12 by blast
+ − 3501
+ − 3502
+ − 3503
lemma contains61a_2:
+ − 3504
assumes "\<exists>r\<in>AALTs_subs (AALTs bs rs). r >> bs2"
+ − 3505
shows "bsimp_AALTs bs rs >> bs2"
+ − 3506
using assms
+ − 3507
apply(induct rs arbitrary: bs2 bs)
+ − 3508
apply(auto)
+ − 3509
apply (simp add: H1 TEMPLATE_contains61a)
+ − 3510
by (metis append_Cons append_Nil contains50 f_cont2)
+ − 3511
+ − 3512
lemma contains_equiv_def2:
+ − 3513
shows " (AALTs bs as >> bs@bs1) \<longleftrightarrow> (\<exists>a\<in>(\<Union> (AALTs_subs ` set as)). a >> bs1)"
+ − 3514
by (metis H1 H3 UN_E UN_I contains0 contains49 contains59 contains60)
+ − 3515
+ − 3516
lemma contains_equiv_def:
+ − 3517
shows "(AALTs bs as >> bs@bs1) \<longleftrightarrow> (\<exists>a\<in>set as. a >> bs1)"
+ − 3518
by (meson contains0 contains49 contains59 contains60)
+ − 3519
+ − 3520
lemma map_fuse2:
+ − 3521
shows "map (bder c) (map (fuse bs) as) = map (fuse bs) (map (bder c) as)"
+ − 3522
by (simp add: map_bder_fuse)
+ − 3523
+ − 3524
lemma map_fuse3:
+ − 3525
shows "map (\<lambda>a. bders a s) (map (fuse bs) as) = map (fuse bs) (map (\<lambda>a. bders a s) as)"
+ − 3526
apply(induct s arbitrary: bs as rule: rev_induct)
+ − 3527
apply(auto simp add: bders_append map_fuse2)
+ − 3528
using bder_fuse by blast
+ − 3529
+ − 3530
lemma bders_AALTs:
+ − 3531
shows "bders (AALTs bs2 as) s = AALTs bs2 (map (\<lambda>a. bders a s) as)"
+ − 3532
apply(induct s arbitrary: bs2 as rule: rev_induct)
+ − 3533
apply(auto simp add: bders_append)
+ − 3534
done
+ − 3535
+ − 3536
lemma bders_AALTs_contains:
+ − 3537
shows "bders (AALTs bs2 as) s >> bs2 @ bs \<longleftrightarrow>
+ − 3538
AALTs bs2 (map (\<lambda>a. bders a s) as) >> bs2 @ bs"
+ − 3539
apply(induct s arbitrary: bs bs2 as)
+ − 3540
apply(auto)[1]
+ − 3541
apply(simp)
+ − 3542
by (smt comp_apply map_eq_conv)
+ − 3543
+ − 3544
+ − 3545
lemma derc_alt00_Urb:
+ − 3546
shows "bder c (bsimp_AALTs bs2 (flts [bsimp a])) >> bs2 @ bs \<longleftrightarrow>
+ − 3547
fuse bs2 (bder c (bsimp a)) >> bs2 @ bs"
+ − 3548
apply(case_tac "bsimp a")
+ − 3549
apply(auto)
+ − 3550
apply(subst (asm) bder_bsimp_AALTs)
+ − 3551
apply(subst (asm) map_fuse2)
+ − 3552
using contains60 contains61 contains63 apply blast
+ − 3553
by (metis bder_bsimp_AALTs contains51c map_bder_fuse map_map)
+ − 3554
+ − 3555
lemma ders_alt00_Urb:
+ − 3556
shows "bders (bsimp_AALTs bs2 (flts [bsimp a])) s >> bs2 @ bs \<longleftrightarrow>
+ − 3557
fuse bs2 (bders (bsimp a) s) >> bs2 @ bs"
+ − 3558
apply(case_tac "bsimp a")
+ − 3559
apply (simp add: bders_AZERO(1))
+ − 3560
using bders_fuse bsimp_AALTs.simps(2) flts.simps(1) flts.simps(4) apply presburger
+ − 3561
using bders_fuse bsimp_AALTs.simps(2) flts.simps(1) flts.simps(5) apply presburger
+ − 3562
using bders_fuse bsimp_AALTs.simps(2) flts.simps(1) flts.simps(6) apply presburger
+ − 3563
prefer 2
+ − 3564
using bders_fuse bsimp_AALTs.simps(2) flts.simps(1) flts.simps(7) apply presburger
+ − 3565
apply(auto simp add: bders_bsimp_AALTs)
+ − 3566
apply(drule contains61)
+ − 3567
apply(auto simp add: bders_AALTs)
+ − 3568
apply(rule contains63)
+ − 3569
apply(rule contains60)
+ − 3570
apply(auto)
+ − 3571
using bders_fuse apply auto[1]
+ − 3572
by (metis contains51c map_fuse3 map_map)
+ − 3573
+ − 3574
lemma derc_alt00_Urb2a:
+ − 3575
shows "bder c (bsimp_AALTs bs2 (flts [bsimp a])) >> bs2 @ bs \<longleftrightarrow>
+ − 3576
bder c (bsimp a) >> bs"
+ − 3577
using contains0 contains49 derc_alt00_Urb by blast
+ − 3578
+ − 3579
+ − 3580
lemma derc_alt00_Urb2:
+ − 3581
assumes "fuse bs2 (bder c (bsimp a)) >> bs2 @ bs" "a \<in> set as"
+ − 3582
shows "bder c (bsimp_AALTs bs2 (flts (map bsimp as))) >> bs2 @ bs"
+ − 3583
using assms
+ − 3584
apply(subgoal_tac "\<exists>list1 list2. as = list1 @ [a] @ list2")
+ − 3585
prefer 2
+ − 3586
using split_list_last apply fastforce
+ − 3587
apply(erule exE)+
+ − 3588
apply(simp add: flts_append del: append.simps)
+ − 3589
using bder_bsimp_AALTs contains50 contains51b derc_alt00_Urb by auto
+ − 3590
+ − 3591
lemma ders_alt00_Urb2:
+ − 3592
assumes "fuse bs2 (bders (bsimp a) s) >> bs2 @ bs" "a \<in> set as"
+ − 3593
shows "bders (bsimp_AALTs bs2 (flts (map bsimp as))) s >> bs2 @ bs"
+ − 3594
using assms
+ − 3595
apply(subgoal_tac "\<exists>list1 list2. as = list1 @ [a] @ list2")
+ − 3596
prefer 2
+ − 3597
using split_list_last apply fastforce
+ − 3598
apply(erule exE)+
+ − 3599
apply(simp add: flts_append del: append.simps)
+ − 3600
apply(simp add: bders_bsimp_AALTs)
+ − 3601
apply(rule contains50)
+ − 3602
apply(rule contains51b)
+ − 3603
using bders_bsimp_AALTs ders_alt00_Urb by auto
+ − 3604
+ − 3605
+ − 3606
lemma derc_alt2:
+ − 3607
assumes "bder c (AALTs bs2 as) >> bs2 @ bs"
+ − 3608
and "\<forall>a \<in> set as. ((bder c a >> bs) \<longrightarrow> (bder c (bsimp a) >> bs))"
+ − 3609
shows "bder c (bsimp (AALTs bs2 as)) >> bs2 @ bs"
+ − 3610
using assms
+ − 3611
apply -
+ − 3612
apply(simp)
+ − 3613
apply(subst (asm) contains_equiv_def)
+ − 3614
apply(simp)
+ − 3615
apply(erule bexE)
+ − 3616
using contains0 derc_alt00_Urb2 by blast
+ − 3617
+ − 3618
+ − 3619
+ − 3620
lemma ders_alt2:
+ − 3621
assumes "bders (AALTs bs2 as) s >> bs2 @ bs"
+ − 3622
and "\<forall>a \<in> set as. ((bders a s >> bs) \<longrightarrow> (bders (bsimp a) s >> bs))"
+ − 3623
shows "bders (bsimp (AALTs bs2 as)) s >> bs2 @ bs"
+ − 3624
using assms
+ − 3625
apply -
+ − 3626
apply(simp add: bders_AALTs)
+ − 3627
thm contains_equiv_def
+ − 3628
apply(subst (asm) contains_equiv_def)
+ − 3629
apply(simp)
+ − 3630
apply(erule bexE)
+ − 3631
using contains0 ders_alt00_Urb2 by blast
+ − 3632
+ − 3633
+ − 3634
+ − 3635
+ − 3636
lemma bder_simp_contains:
+ − 3637
assumes "bder c a >> bs"
+ − 3638
shows "bder c (bsimp a) >> bs"
+ − 3639
using assms
+ − 3640
apply(induct a arbitrary: c bs)
+ − 3641
apply(auto elim: contains.cases)
+ − 3642
apply(case_tac "bnullable a1")
+ − 3643
apply(simp)
+ − 3644
prefer 2
+ − 3645
apply(simp)
+ − 3646
apply(erule contains.cases)
+ − 3647
apply(auto)
+ − 3648
apply(case_tac "(bsimp a1) = AZERO")
+ − 3649
apply(simp)
+ − 3650
apply (metis append_Nil2 contains0 contains49 fuse.simps(1))
+ − 3651
apply(case_tac "(bsimp a2a) = AZERO")
+ − 3652
apply(simp)
+ − 3653
apply (metis bder.simps(1) bsimp.simps(1) bsimp_ASEQ0 contains.intros(3) contains55)
+ − 3654
apply(case_tac "\<exists>bsX. (bsimp a1) = AONE bsX")
+ − 3655
apply(auto)[1]
+ − 3656
using b3 apply fastforce
+ − 3657
apply(subst bsimp_ASEQ1)
+ − 3658
apply(auto)[3]
+ − 3659
apply(simp)
+ − 3660
apply(subgoal_tac "\<not> bnullable (bsimp a1)")
+ − 3661
prefer 2
+ − 3662
using b3 apply blast
+ − 3663
apply(simp)
+ − 3664
apply (simp add: contains.intros(3) contains55)
+ − 3665
(* SEQ nullable case *)
+ − 3666
apply(erule contains.cases)
+ − 3667
apply(auto)
+ − 3668
apply(erule contains.cases)
+ − 3669
apply(auto)
+ − 3670
apply(case_tac "(bsimp a1) = AZERO")
+ − 3671
apply(simp)
+ − 3672
apply (metis append_Nil2 contains0 contains49 fuse.simps(1))
+ − 3673
apply(case_tac "(bsimp a2a) = AZERO")
+ − 3674
apply(simp)
+ − 3675
apply (metis bder.simps(1) bsimp.simps(1) bsimp_ASEQ0 contains.intros(3) contains55)
+ − 3676
apply(case_tac "\<exists>bsX. (bsimp a1) = AONE bsX")
+ − 3677
apply(auto)[1]
+ − 3678
using contains.simps apply blast
+ − 3679
apply(subst bsimp_ASEQ1)
+ − 3680
apply(auto)[3]
+ − 3681
apply(simp)
+ − 3682
apply(subgoal_tac "bnullable (bsimp a1)")
+ − 3683
prefer 2
+ − 3684
using b3 apply blast
+ − 3685
apply(simp)
+ − 3686
apply (metis contains.intros(3) contains.intros(4) contains55 self_append_conv2)
+ − 3687
apply(erule contains.cases)
+ − 3688
apply(auto)
+ − 3689
apply(case_tac "(bsimp a1) = AZERO")
+ − 3690
apply(simp)
+ − 3691
using b3 apply force
+ − 3692
apply(case_tac "(bsimp a2) = AZERO")
+ − 3693
apply(simp)
+ − 3694
apply (metis bder.simps(1) bsimp_ASEQ0 bsimp_ASEQ_fuse contains0 contains49 f_cont1)
+ − 3695
apply(case_tac "\<exists>bsX. (bsimp a1) = AONE bsX")
+ − 3696
apply(auto)[1]
+ − 3697
apply (metis append_assoc bder_fuse bmkeps.simps(1) bmkeps_simp bsimp_ASEQ2 contains0 contains49 f_cont1)
+ − 3698
apply(subst bsimp_ASEQ1)
+ − 3699
apply(auto)[3]
+ − 3700
apply(simp)
+ − 3701
apply(subgoal_tac "bnullable (bsimp a1)")
+ − 3702
prefer 2
+ − 3703
using b3 apply blast
+ − 3704
apply(simp)
+ − 3705
apply (metis bmkeps_simp contains.intros(4) contains.intros(5) contains0 contains49 f_cont1)
+ − 3706
apply(erule contains.cases)
+ − 3707
apply(auto)
+ − 3708
(* ALT case *)
+ − 3709
apply(subgoal_tac "\<exists>bsX. bs = x1 @ bsX")
+ − 3710
prefer 2
+ − 3711
using contains59 f_cont1 apply blast
+ − 3712
apply(auto)
+ − 3713
apply(rule derc_alt2[simplified])
+ − 3714
apply(simp)
+ − 3715
by blast
+ − 3716
+ − 3717
+ − 3718
+ − 3719
lemma bder_simp_containsA:
+ − 3720
assumes "bder c a >> bs"
+ − 3721
shows "bsimp (bder c (bsimp a)) >> bs"
+ − 3722
using assms
+ − 3723
by (simp add: bder_simp_contains contains55)
+ − 3724
+ − 3725
lemma bder_simp_containsB:
+ − 3726
assumes "bsimp (bder c a) >> bs"
+ − 3727
shows "bder c (bsimp a) >> bs"
+ − 3728
using assms
+ − 3729
by (simp add: PPP1_eq bder_simp_contains)
+ − 3730
+ − 3731
lemma bder_simp_contains_IFF:
+ − 3732
assumes "good a"
+ − 3733
shows "bsimp (bder c a) >> bs \<longleftrightarrow> bder c (bsimp a) >> bs"
+ − 3734
using assms
+ − 3735
by (simp add: PPP1_eq test2)
+ − 3736
+ − 3737
+ − 3738
lemma ders_seq:
+ − 3739
assumes "bders (ASEQ bs a1 a2) s >> bs @ bs2"
+ − 3740
and "\<And>s bs. bders a1 s >> bs \<Longrightarrow> bders (bsimp a1) s >> bs"
+ − 3741
"\<And>s bs. bders a2 s >> bs \<Longrightarrow> bders (bsimp a2) s >> bs"
+ − 3742
shows "bders (ASEQ bs (bsimp a1) (bsimp a2)) s >> bs @ bs2"
+ − 3743
using assms(1)
+ − 3744
apply(induct s arbitrary: a1 a2 bs bs2 rule: rev_induct)
+ − 3745
apply(auto)[1]
+ − 3746
thm CT1_SEQ PPP1_eq
+ − 3747
apply (metis CT1_SEQ PPP1_eq)
+ − 3748
apply(auto simp add: bders_append)
+ − 3749
apply(drule bder_simp_contains)
+ − 3750
oops
+ − 3751
+ − 3752
+ − 3753
lemma bders_simp_contains:
+ − 3754
assumes "bders a s >> bs"
+ − 3755
shows "bders (bsimp a) s >> bs"
+ − 3756
using assms
+ − 3757
apply(induct a arbitrary: s bs)
+ − 3758
apply(auto elim: contains.cases)[4]
+ − 3759
prefer 2
+ − 3760
apply(subgoal_tac "\<exists>bsX. bs = x1 @ bsX")
+ − 3761
prefer 2
+ − 3762
apply (metis bders_AALTs contains59 f_cont1)
+ − 3763
apply(clarify)
+ − 3764
apply(rule ders_alt2)
+ − 3765
apply(assumption)
+ − 3766
apply(auto)[1]
+ − 3767
prefer 2
+ − 3768
apply simp
+ − 3769
(* SEQ case *)
+ − 3770
apply(case_tac "bsimp a1 = AZERO")
+ − 3771
apply(simp)
+ − 3772
apply (metis LLLL(1) bders_AZERO(1) bsimp.simps(1) bsimp.simps(3) bsimp_ASEQ.simps(1) contains55 ders_correctness erase_bders good.simps(1) good1a xxx_bder2)
+ − 3773
apply(case_tac "bsimp a2 = AZERO")
+ − 3774
apply(simp)
+ − 3775
apply (metis LLLL(1) bders_AZERO(1) bsimp.simps(1) bsimp.simps(3) bsimp_ASEQ0 contains55 ders_correctness erase_bders good.simps(1) good1a xxx_bder2)
+ − 3776
apply(case_tac "\<exists>bsX. bsimp a1 = AONE bsX")
+ − 3777
apply(auto)
+ − 3778
apply(subst bsimp_ASEQ2)
+ − 3779
apply(case_tac s)
+ − 3780
apply(simp)
+ − 3781
apply (metis b1 bsimp.simps(1) contains55)
+ − 3782
apply(simp)
+ − 3783
apply(subgoal_tac "bnullable a1")
+ − 3784
prefer 2
+ − 3785
using b3 apply fastforce
+ − 3786
apply(auto)
+ − 3787
apply(subst (asm) bders_AALTs)
+ − 3788
apply(erule contains.cases)
+ − 3789
apply(auto)
+ − 3790
prefer 2
+ − 3791
apply(erule contains.cases)
+ − 3792
apply(auto)
+ − 3793
apply(simp add: fuse_append)
+ − 3794
apply(simp add: bder_fuse bders_fuse)
+ − 3795
apply (metis bders.simps(2) bmkeps.simps(1) bmkeps_simp contains0 contains49 f_cont1)
+ − 3796
using contains_equiv_def apply auto[1]
+ − 3797
apply(simp add: bder_fuse bders_fuse fuse_append)
+ − 3798
apply(rule contains0)
+ − 3799
oops
+ − 3800
+ − 3801
+ − 3802
lemma T0:
+ − 3803
assumes "s = []"
+ − 3804
shows "bders (bsimp r) s >> bs \<longleftrightarrow> bders r s >> bs"
+ − 3805
using assms
+ − 3806
by (simp add: PPP1_eq test2)
+ − 3807
+ − 3808
lemma T1:
+ − 3809
assumes "s = [a]" "bders r s >> bs"
+ − 3810
shows "bders (bsimp r) s >> bs"
+ − 3811
using assms
+ − 3812
apply(simp)
+ − 3813
by (simp add: bder_simp_contains)
+ − 3814
+ − 3815
lemma TX:
+ − 3816
assumes "\<Turnstile> v : ders s (erase r)" "\<Turnstile> v : ders s (erase (bsimp r))"
+ − 3817
shows "bders r s >> FC r s v \<longleftrightarrow> bders (bsimp r) s >> FC (bsimp r) s v"
+ − 3818
using FC_def contains7b
+ − 3819
using assms by metis
+ − 3820
+ − 3821
lemma mkeps1:
+ − 3822
assumes "s \<in> L (erase r)"
+ − 3823
shows "\<Turnstile> mkeps (ders s (erase r)) : ders s (erase r)"
+ − 3824
using assms
+ − 3825
by (meson lexer_correct_None lexer_flex mkeps_nullable)
+ − 3826
+ − 3827
lemma mkeps2:
+ − 3828
assumes "s \<in> L (erase r)"
+ − 3829
shows "\<Turnstile> mkeps (ders s (erase (bsimp r))) : ders s (erase (bsimp r))"
+ − 3830
using assms
+ − 3831
by (metis LLLL(1) lexer_correct_None lexer_flex mkeps_nullable)
+ − 3832
+ − 3833
thm FC_def FE_def PX_def PV_def
+ − 3834
+ − 3835
+ − 3836
lemma TX2:
+ − 3837
assumes "s \<in> L (erase r)"
+ − 3838
shows "bders r s >> FE r s \<longleftrightarrow> bders (bsimp r) s >> FE (bsimp r) s"
+ − 3839
using assms
+ − 3840
by (simp add: FE_def contains7b mkeps1 mkeps2)
+ − 3841
+ − 3842
lemma TX3:
+ − 3843
assumes "s \<in> L (erase r)"
+ − 3844
shows "bders r s >> FE r s \<longleftrightarrow> bders (bsimp r) s >> FE (bders (bsimp r) s) []"
+ − 3845
using assms
+ − 3846
by (metis FE_PX FE_def L07 LLLL(1) PX_id TX2)
+ − 3847
+ − 3848
find_theorems "FE _ _ = _"
+ − 3849
find_theorems "FC _ _ _ = _"
+ − 3850
find_theorems "(bder _ _ >> _ _ _ _) = _"
+ − 3851
+ − 3852
+ − 3853
(* HERE *)
+ − 3854
+ − 3855
lemma PX:
+ − 3856
assumes "s \<in> L r" "bders (intern r) s >> code (PX r s)"
+ − 3857
shows "bders (bsimp (intern r)) s >> code (PX r s)"
+ − 3858
using assms
+ − 3859
apply(induct s arbitrary: r rule: rev_induct)
+ − 3860
apply(simp)
+ − 3861
apply (simp add: PPP1_eq)
+ − 3862
apply (simp add: bders_append bders_simp_append)
+ − 3863
thm PX_bder_iff PX_bders_iff
+ − 3864
apply(subst (asm) PX_bder_iff)
+ − 3865
apply(assumption)
+ − 3866
apply(subst (asm) (2) PX_bders_iff)
+ − 3867
find_theorems "_ >> code (PX _ _)"
+ − 3868
find_theorems "PX _ _ = _"
+ − 3869
find_theorems "(intern _) >> _"
+ − 3870
apply (simp add: contains55)
+ − 3871
apply (simp add: bders_append bders_simp_append)
+ − 3872
apply (simp add: PPP1_eq)
+ − 3873
find_theorems "(bder _ _ >> _) = _"
+ − 3874
apply(rule contains50)
+ − 3875
+ − 3876
apply(case_tac "bders a xs = AZERO")
+ − 3877
apply(simp)
+ − 3878
apply(subgoal_tac "bders_simp a xs = AZERO")
+ − 3879
prefer 2
+ − 3880
apply (metis L_bders_simp XXX4a_good_cons bders.simps(1) bders_simp.simps(1) bsimp.simps(3) good.simps(1) good1a test2 xxx_bder2)
+ − 3881
apply(simp)
+ − 3882
apply(case_tac xs)
+ − 3883
apply(simp)
+ − 3884
apply (simp add: PPP1_eq)
+ − 3885
apply(simp)
+ − 3886
apply(subgoal_tac "good (bders_simp a (aa # list)) \<or> (bders_simp a (aa # list) = AZERO)")
+ − 3887
apply(auto)
+ − 3888
apply(subst (asm) bder_simp_contains_IFF)
+ − 3889
apply(simp)
+ − 3890
+ − 3891
(* TOBE PROVED *)
+ − 3892
lemma
+ − 3893
assumes "s \<in> L (erase r)"
+ − 3894
shows "bders_simp r s >> bs \<longleftrightarrow> bders r s >> bs"
+ − 3895
using assms
+ − 3896
apply(induct s arbitrary: r bs)
+ − 3897
apply(simp)
+ − 3898
apply(simp add: bders_append bders_simp_append)
+ − 3899
apply(rule iffI)
+ − 3900
apply(drule_tac x="bsimp (bder a r)" in meta_spec)
+ − 3901
apply(drule_tac x="bs" in meta_spec)
+ − 3902
apply(drule meta_mp)
+ − 3903
using L_bsimp_erase lexer_correct_None apply fastforce
+ − 3904
apply(simp)
+ − 3905
+ − 3906
+ − 3907
prefer 2
+ − 3908
+ − 3909
+ − 3910
oops
+ − 3911
+ − 3912
+ − 3913
lemma
+ − 3914
assumes "s \<in> L r"
+ − 3915
shows "(bders_simp (intern r) s >> code (PX r s)) \<longleftrightarrow> ((intern r) >> code (PX r s))"
+ − 3916
using assms
+ − 3917
apply(induct s arbitrary: r rule: rev_induct)
+ − 3918
apply(simp)
+ − 3919
apply(simp add: bders_simp_append)
+ − 3920
apply(simp add: PPP1_eq)
+ − 3921
+ − 3922
+ − 3923
find_theorems "retrieve (bders _ _) _"
+ − 3924
find_theorems "_ >> retrieve _ _"
+ − 3925
find_theorems "bsimp _ >> _"
+ − 3926
oops
+ − 3927
+ − 3928
+ − 3929
lemma PX4a:
+ − 3930
assumes "(s1 @ s2) \<in> L r"
+ − 3931
shows "bders (intern r) (s1 @ s2) >> code (PV r s1 (PX (ders s1 r) s2))"
+ − 3932
using PX4[OF assms]
+ − 3933
apply(simp add: PX_append)
+ − 3934
done
+ − 3935
+ − 3936
lemma PV5:
+ − 3937
assumes "s2 \<in> (ders s1 r) \<rightarrow> v"
+ − 3938
shows "bders (intern r) (s1 @ s2) >> code (PV r s1 v)"
+ − 3939
by (simp add: PPP0_isar PV_def Posix_flex assms)
+ − 3940
+ − 3941
lemma PV6:
+ − 3942
assumes "s2 \<in> (ders s1 r) \<rightarrow> v"
+ − 3943
shows "bders (bders (intern r) s1) s2 >> code (PV r s1 v)"
+ − 3944
using PV5 assms bders_append by auto
+ − 3945
+ − 3946
find_theorems "retrieve (bders _ _) _"
+ − 3947
find_theorems "_ >> retrieve _ _"
+ − 3948
find_theorems "bder _ _ >> _"
+ − 3949
+ − 3950
+ − 3951
lemma OO0_PX:
+ − 3952
assumes "s \<in> L r"
+ − 3953
shows "bders (intern r) s >> code (PX r s)"
+ − 3954
using assms
+ − 3955
by (simp add: PX3)
+ − 3956
+ − 3957
+ − 3958
lemma OO1:
+ − 3959
assumes "[c] \<in> r \<rightarrow> v"
+ − 3960
shows "bder c (intern r) >> code v"
+ − 3961
using assms
+ − 3962
using PPP0_isar by force
+ − 3963
+ − 3964
lemma OO1a:
+ − 3965
assumes "[c] \<in> L r"
+ − 3966
shows "bder c (intern r) >> code (PX r [c])"
+ − 3967
using assms unfolding PX_def PV_def
+ − 3968
using contains70 by fastforce
+ − 3969
+ − 3970
lemma OO12:
+ − 3971
assumes "[c1, c2] \<in> L r"
+ − 3972
shows "bders (intern r) [c1, c2] >> code (PX r [c1, c2])"
+ − 3973
using assms
+ − 3974
using PX_def PV_def contains70 by presburger
+ − 3975
+ − 3976
lemma OO2:
+ − 3977
assumes "[c] \<in> L r"
+ − 3978
shows "bders_simp (intern r) [c] >> code (PX r [c])"
+ − 3979
using assms
+ − 3980
using OO1a Posix1(1) contains55 by auto
+ − 3981
+ − 3982
+ − 3983
thm L07XX PPP0b erase_intern
+ − 3984
+ − 3985
find_theorems "retrieve (bders _ _) _"
+ − 3986
find_theorems "_ >> retrieve _ _"
+ − 3987
find_theorems "bder _ _ >> _"
+ − 3988
+ − 3989
+ − 3990
lemma PPP3:
+ − 3991
assumes "\<Turnstile> v : ders s (erase a)"
+ − 3992
shows "bders a s >> retrieve a (flex (erase a) id s v)"
+ − 3993
using LA[OF assms] contains6 erase_bders assms by metis
+ − 3994
+ − 3995
+ − 3996
find_theorems "bder _ _ >> _"
+ − 3997
+ − 3998
+ − 3999
lemma
+ − 4000
fixes n :: nat
+ − 4001
shows "(\<Sum>i \<in> {0..n}. i) = n * (n + 1) div 2"
+ − 4002
apply(induct n)
+ − 4003
apply(simp)
+ − 4004
apply(simp)
+ − 4005
done
+ − 4006
+ − 4007
lemma COUNTEREXAMPLE:
+ − 4008
assumes "r = AALTs [S] [ASEQ [S] (AALTs [S] [AONE [S], ACHAR [S] c]) (ACHAR [S] c)]"
+ − 4009
shows "bsimp (bder c (bsimp r)) = bsimp (bder c r)"
+ − 4010
apply(simp_all add: assms)
+ − 4011
oops
+ − 4012
+ − 4013
lemma COUNTEREXAMPLE:
+ − 4014
assumes "r = AALTs [S] [ASEQ [S] (AALTs [S] [AONE [S], ACHAR [S] c]) (ACHAR [S] c)]"
+ − 4015
shows "bsimp r = r"
+ − 4016
apply(simp_all add: assms)
+ − 4017
oops
+ − 4018
+ − 4019
lemma COUNTEREXAMPLE:
+ − 4020
assumes "r = AALTs [S] [ASEQ [S] (AALTs [S] [AONE [S], ACHAR [S] c]) (ACHAR [S] c)]"
+ − 4021
shows "bsimp r = XXX"
+ − 4022
and "bder c r = XXX"
+ − 4023
and "bder c (bsimp r) = XXX"
+ − 4024
and "bsimp (bder c (bsimp r)) = XXX"
+ − 4025
and "bsimp (bder c r) = XXX"
+ − 4026
apply(simp_all add: assms)
+ − 4027
oops
+ − 4028
+ − 4029
lemma COUNTEREXAMPLE_contains1:
+ − 4030
assumes "r = AALTs [S] [ASEQ [S] (AALTs [S] [AONE [S], ACHAR [S] c]) (ACHAR [S] c)]"
+ − 4031
and "bsimp (bder c r) >> bs"
+ − 4032
shows "bsimp (bder c (bsimp r)) >> bs"
+ − 4033
using assms
+ − 4034
apply(auto elim!: contains.cases)
+ − 4035
apply(rule Etrans)
+ − 4036
apply(rule contains.intros)
+ − 4037
apply(rule contains.intros)
+ − 4038
apply(simp)
+ − 4039
apply(rule Etrans)
+ − 4040
apply(rule contains.intros)
+ − 4041
apply(rule contains.intros)
+ − 4042
apply(simp)
+ − 4043
done
+ − 4044
+ − 4045
lemma COUNTEREXAMPLE_contains2:
+ − 4046
assumes "r = AALTs [S] [ASEQ [S] (AALTs [S] [AONE [S], ACHAR [S] c]) (ACHAR [S] c)]"
+ − 4047
and "bsimp (bder c (bsimp r)) >> bs"
+ − 4048
shows "bsimp (bder c r) >> bs"
+ − 4049
using assms
+ − 4050
apply(auto elim!: contains.cases)
+ − 4051
apply(rule Etrans)
+ − 4052
apply(rule contains.intros)
+ − 4053
apply(rule contains.intros)
+ − 4054
apply(simp)
+ − 4055
apply(rule Etrans)
+ − 4056
apply(rule contains.intros)
+ − 4057
apply(rule contains.intros)
+ − 4058
apply(simp)
+ − 4059
done
+ − 4060
+ − 4061
+ − 4062
end