thys3/BlexerSimp.thy
author Chengsong
Sat, 16 Jul 2022 18:34:46 +0100
changeset 568 7a579f5533f8
parent 564 3cbcd7cda0a9
child 575 3178f0e948ac
permissions -rw-r--r--
more chapter2 modifications
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
544
Chengsong
parents:
diff changeset
     1
theory BlexerSimp
Chengsong
parents:
diff changeset
     2
  imports Blexer 
Chengsong
parents:
diff changeset
     3
begin
Chengsong
parents:
diff changeset
     4
Chengsong
parents:
diff changeset
     5
Chengsong
parents:
diff changeset
     6
fun flts :: "arexp list \<Rightarrow> arexp list"
Chengsong
parents:
diff changeset
     7
  where 
Chengsong
parents:
diff changeset
     8
  "flts [] = []"
Chengsong
parents:
diff changeset
     9
| "flts (AZERO # rs) = flts rs"
Chengsong
parents:
diff changeset
    10
| "flts ((AALTs bs  rs1) # rs) = (map (fuse bs) rs1) @ flts rs"
Chengsong
parents:
diff changeset
    11
| "flts (r1 # rs) = r1 # flts rs"
Chengsong
parents:
diff changeset
    12
Chengsong
parents:
diff changeset
    13
Chengsong
parents:
diff changeset
    14
Chengsong
parents:
diff changeset
    15
fun bsimp_ASEQ :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp \<Rightarrow> arexp"
Chengsong
parents:
diff changeset
    16
  where
Chengsong
parents:
diff changeset
    17
  "bsimp_ASEQ _ AZERO _ = AZERO"
Chengsong
parents:
diff changeset
    18
| "bsimp_ASEQ _ _ AZERO = AZERO"
Chengsong
parents:
diff changeset
    19
| "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2"
Chengsong
parents:
diff changeset
    20
| "bsimp_ASEQ bs1 r1 r2 = ASEQ  bs1 r1 r2"
Chengsong
parents:
diff changeset
    21
Chengsong
parents:
diff changeset
    22
lemma bsimp_ASEQ0[simp]:
Chengsong
parents:
diff changeset
    23
  shows "bsimp_ASEQ bs r1 AZERO = AZERO"
Chengsong
parents:
diff changeset
    24
  by (case_tac r1)(simp_all)
Chengsong
parents:
diff changeset
    25
Chengsong
parents:
diff changeset
    26
lemma bsimp_ASEQ1:
Chengsong
parents:
diff changeset
    27
  assumes "r1 \<noteq> AZERO" "r2 \<noteq> AZERO" "\<nexists>bs. r1 = AONE bs"
Chengsong
parents:
diff changeset
    28
  shows "bsimp_ASEQ bs r1 r2 = ASEQ bs r1 r2"
Chengsong
parents:
diff changeset
    29
  using assms
Chengsong
parents:
diff changeset
    30
  apply(induct bs r1 r2 rule: bsimp_ASEQ.induct)
Chengsong
parents:
diff changeset
    31
  apply(auto)
Chengsong
parents:
diff changeset
    32
  done
Chengsong
parents:
diff changeset
    33
Chengsong
parents:
diff changeset
    34
lemma bsimp_ASEQ2[simp]:
Chengsong
parents:
diff changeset
    35
  shows "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2"
Chengsong
parents:
diff changeset
    36
  by (case_tac r2) (simp_all)
Chengsong
parents:
diff changeset
    37
Chengsong
parents:
diff changeset
    38
Chengsong
parents:
diff changeset
    39
fun bsimp_AALTs :: "bit list \<Rightarrow> arexp list \<Rightarrow> arexp"
Chengsong
parents:
diff changeset
    40
  where
Chengsong
parents:
diff changeset
    41
  "bsimp_AALTs _ [] = AZERO"
Chengsong
parents:
diff changeset
    42
| "bsimp_AALTs bs1 [r] = fuse bs1 r"
Chengsong
parents:
diff changeset
    43
| "bsimp_AALTs bs1 rs = AALTs bs1 rs"
Chengsong
parents:
diff changeset
    44
Chengsong
parents:
diff changeset
    45
lemma bmkeps_fuse:
Chengsong
parents:
diff changeset
    46
  assumes "bnullable r"
Chengsong
parents:
diff changeset
    47
  shows "bmkeps (fuse bs r) = bs @ bmkeps r"
Chengsong
parents:
diff changeset
    48
  using assms
Chengsong
parents:
diff changeset
    49
  by (induct r rule: bnullable.induct) (auto)
Chengsong
parents:
diff changeset
    50
Chengsong
parents:
diff changeset
    51
lemma bmkepss_fuse: 
Chengsong
parents:
diff changeset
    52
  assumes "bnullables rs"
Chengsong
parents:
diff changeset
    53
  shows "bmkepss (map (fuse bs) rs) = bs @ bmkepss rs"
Chengsong
parents:
diff changeset
    54
  using assms
Chengsong
parents:
diff changeset
    55
  apply(induct rs arbitrary: bs)
Chengsong
parents:
diff changeset
    56
  apply(auto simp add: bmkeps_fuse bnullable_fuse)
Chengsong
parents:
diff changeset
    57
  done
Chengsong
parents:
diff changeset
    58
Chengsong
parents:
diff changeset
    59
lemma bder_fuse:
Chengsong
parents:
diff changeset
    60
  shows "bder c (fuse bs a) = fuse bs  (bder c a)"
Chengsong
parents:
diff changeset
    61
  apply(induct a arbitrary: bs c)
Chengsong
parents:
diff changeset
    62
  apply(simp_all)
Chengsong
parents:
diff changeset
    63
  done
Chengsong
parents:
diff changeset
    64
Chengsong
parents:
diff changeset
    65
Chengsong
parents:
diff changeset
    66
Chengsong
parents:
diff changeset
    67
550
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    68
fun ABIncludedByC :: "'a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool" where
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    69
  "ABIncludedByC a b c f subseteqPred = subseteqPred (f a b) c"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    70
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    71
fun furtherSEQ :: "rexp \<Rightarrow> rexp \<Rightarrow> rexp list" and 
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    72
   turnIntoTerms :: "rexp \<Rightarrow> rexp list "
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    73
   where
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    74
  "furtherSEQ ONE r2 =  turnIntoTerms r2 "
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    75
| "furtherSEQ r11 r2 = [SEQ r11 r2]"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    76
| "turnIntoTerms (SEQ ONE r2) =  turnIntoTerms r2"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    77
| "turnIntoTerms (SEQ r1 r2) = concat (map (\<lambda>r11. furtherSEQ r11 r2) (turnIntoTerms r1))"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    78
| "turnIntoTerms r = [r]"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    79
564
Chengsong
parents: 558
diff changeset
    80
abbreviation "tint \<equiv> turnIntoTerms"
Chengsong
parents: 558
diff changeset
    81
Chengsong
parents: 558
diff changeset
    82
fun seqFold :: "rexp \<Rightarrow> rexp list \<Rightarrow> rexp" where
Chengsong
parents: 558
diff changeset
    83
  "seqFold acc [] = acc"
Chengsong
parents: 558
diff changeset
    84
| "seqFold ONE (r # rs1) = seqFold r rs1"
Chengsong
parents: 558
diff changeset
    85
| "seqFold acc (r # rs1) = seqFold (SEQ acc r) rs1"
Chengsong
parents: 558
diff changeset
    86
550
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    87
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    88
fun attachCtx :: "arexp \<Rightarrow> rexp list \<Rightarrow> rexp set" where
564
Chengsong
parents: 558
diff changeset
    89
  "attachCtx r ctx = set (turnIntoTerms (seqFold (erase r) ctx))"
550
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    90
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    91
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    92
fun rs1_subseteq_rs2 :: "rexp set \<Rightarrow> rexp set \<Rightarrow> bool" where
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    93
  "rs1_subseteq_rs2 rs1 rs2 = (rs1 \<subseteq> rs2)"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    94
564
Chengsong
parents: 558
diff changeset
    95
550
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    96
fun notZero :: "arexp \<Rightarrow> bool" where
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    97
  "notZero AZERO = True"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    98
| "notZero _ = False"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
    99
564
Chengsong
parents: 558
diff changeset
   100
558
Chengsong
parents: 557
diff changeset
   101
fun tset :: "arexp list \<Rightarrow> rexp set" where
Chengsong
parents: 557
diff changeset
   102
  "tset rs = set (concat (map turnIntoTerms (map erase rs)))"
Chengsong
parents: 557
diff changeset
   103
Chengsong
parents: 557
diff changeset
   104
550
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   105
fun prune6 :: "rexp set \<Rightarrow> arexp \<Rightarrow> rexp list \<Rightarrow> arexp" where
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   106
  "prune6 acc a ctx = (if (ABIncludedByC a ctx acc attachCtx rs1_subseteq_rs2) then AZERO else 
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   107
                        (case a of (ASEQ bs r1 r2) \<Rightarrow> bsimp_ASEQ bs (prune6 acc r1 (erase r2 # ctx)) r2
564
Chengsong
parents: 558
diff changeset
   108
                                 | AALTs bs rs0 \<Rightarrow> bsimp_AALTs bs (filter notZero (map (\<lambda>r.(prune6 acc r ctx)) rs0))
Chengsong
parents: 558
diff changeset
   109
                                 | r \<Rightarrow> r
Chengsong
parents: 558
diff changeset
   110
                        )
Chengsong
parents: 558
diff changeset
   111
                      )
Chengsong
parents: 558
diff changeset
   112
            "
550
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   113
558
Chengsong
parents: 557
diff changeset
   114
abbreviation
564
Chengsong
parents: 558
diff changeset
   115
  "p6 acc r \<equiv> prune6 (tset acc) r Nil"
550
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   116
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   117
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   118
fun dB6 :: "arexp list \<Rightarrow> rexp set \<Rightarrow> arexp list" where
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   119
  "dB6 [] acc = []"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   120
| "dB6 (a # as) acc = (if (erase a \<in> acc) then (dB6 as acc) 
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   121
                       else (let pruned = prune6 acc a [] in 
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   122
                              (case pruned of
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   123
                                 AZERO \<Rightarrow> dB6 as acc
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   124
                               |xPrime \<Rightarrow> xPrime # (dB6 as ( (set (turnIntoTerms (erase pruned))) \<union> acc)  ) ) ))   "
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   125
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   126
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   127
fun bsimpStrong6 :: "arexp \<Rightarrow> arexp" 
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   128
  where
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   129
  "bsimpStrong6 (ASEQ bs1 r1 r2) = bsimp_ASEQ bs1 (bsimpStrong6 r1) (bsimpStrong6 r2)"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   130
| "bsimpStrong6 (AALTs bs1 rs) = bsimp_AALTs bs1 (dB6 (flts (map bsimpStrong6 rs)) {}) "
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   131
| "bsimpStrong6 r = r"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   132
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   133
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   134
fun 
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   135
  bdersStrong6 :: "arexp \<Rightarrow> string \<Rightarrow> arexp"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   136
where
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   137
  "bdersStrong6 r [] = r"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   138
| "bdersStrong6 r (c # s) = bdersStrong6 (bsimpStrong6 (bder c r)) s"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   139
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   140
definition blexerStrong where
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   141
 "blexerStrong r s \<equiv> if bnullable (bdersStrong6 (intern r) s) then 
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   142
                    decode (bmkeps (bdersStrong6 (intern r) s)) r else None"
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   143
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   144
9feeb279afdd more deletion of bsimp related
Chengsong
parents: 549
diff changeset
   145
544
Chengsong
parents:
diff changeset
   146
inductive 
Chengsong
parents:
diff changeset
   147
  rrewrite:: "arexp \<Rightarrow> arexp \<Rightarrow> bool" ("_ \<leadsto> _" [99, 99] 99)
Chengsong
parents:
diff changeset
   148
and 
Chengsong
parents:
diff changeset
   149
  srewrite:: "arexp list \<Rightarrow> arexp list \<Rightarrow> bool" (" _ s\<leadsto> _" [100, 100] 100)
Chengsong
parents:
diff changeset
   150
where
Chengsong
parents:
diff changeset
   151
  bs1: "ASEQ bs AZERO r2 \<leadsto> AZERO"
Chengsong
parents:
diff changeset
   152
| bs2: "ASEQ bs r1 AZERO \<leadsto> AZERO"
Chengsong
parents:
diff changeset
   153
| bs3: "ASEQ bs1 (AONE bs2) r \<leadsto> fuse (bs1@bs2) r"
Chengsong
parents:
diff changeset
   154
| bs4: "r1 \<leadsto> r2 \<Longrightarrow> ASEQ bs r1 r3 \<leadsto> ASEQ bs r2 r3"
Chengsong
parents:
diff changeset
   155
| bs5: "r3 \<leadsto> r4 \<Longrightarrow> ASEQ bs r1 r3 \<leadsto> ASEQ bs r1 r4"
Chengsong
parents:
diff changeset
   156
| bs6: "AALTs bs [] \<leadsto> AZERO"
Chengsong
parents:
diff changeset
   157
| bs7: "AALTs bs [r] \<leadsto> fuse bs r"
Chengsong
parents:
diff changeset
   158
| bs10: "rs1 s\<leadsto> rs2 \<Longrightarrow> AALTs bs rs1 \<leadsto> AALTs bs rs2"
Chengsong
parents:
diff changeset
   159
| ss1:  "[] s\<leadsto> []"
Chengsong
parents:
diff changeset
   160
| ss2:  "rs1 s\<leadsto> rs2 \<Longrightarrow> (r # rs1) s\<leadsto> (r # rs2)"
Chengsong
parents:
diff changeset
   161
| ss3:  "r1 \<leadsto> r2 \<Longrightarrow> (r1 # rs) s\<leadsto> (r2 # rs)"
Chengsong
parents:
diff changeset
   162
| ss4:  "(AZERO # rs) s\<leadsto> rs"
Chengsong
parents:
diff changeset
   163
| ss5:  "(AALTs bs1 rs1 # rsb) s\<leadsto> ((map (fuse bs1) rs1) @ rsb)"
Chengsong
parents:
diff changeset
   164
| ss6:  "L (erase a2) \<subseteq> L (erase a1) \<Longrightarrow> (rsa@[a1]@rsb@[a2]@rsc) s\<leadsto> (rsa@[a1]@rsb@rsc)"
564
Chengsong
parents: 558
diff changeset
   165
| ss7:  " (as @ [a] @ as1) s\<leadsto> (as @ [p6 as a] @ as1)"
544
Chengsong
parents:
diff changeset
   166
564
Chengsong
parents: 558
diff changeset
   167
thm tset.simps
544
Chengsong
parents:
diff changeset
   168
Chengsong
parents:
diff changeset
   169
inductive 
Chengsong
parents:
diff changeset
   170
  rrewrites:: "arexp \<Rightarrow> arexp \<Rightarrow> bool" ("_ \<leadsto>* _" [100, 100] 100)
Chengsong
parents:
diff changeset
   171
where 
Chengsong
parents:
diff changeset
   172
  rs1[intro, simp]:"r \<leadsto>* r"
Chengsong
parents:
diff changeset
   173
| rs2[intro]: "\<lbrakk>r1 \<leadsto>* r2; r2 \<leadsto> r3\<rbrakk> \<Longrightarrow> r1 \<leadsto>* r3"
Chengsong
parents:
diff changeset
   174
Chengsong
parents:
diff changeset
   175
inductive 
Chengsong
parents:
diff changeset
   176
  srewrites:: "arexp list \<Rightarrow> arexp list \<Rightarrow> bool" ("_ s\<leadsto>* _" [100, 100] 100)
Chengsong
parents:
diff changeset
   177
where 
Chengsong
parents:
diff changeset
   178
  sss1[intro, simp]:"rs s\<leadsto>* rs"
Chengsong
parents:
diff changeset
   179
| sss2[intro]: "\<lbrakk>rs1 s\<leadsto> rs2; rs2 s\<leadsto>* rs3\<rbrakk> \<Longrightarrow> rs1 s\<leadsto>* rs3"
Chengsong
parents:
diff changeset
   180
Chengsong
parents:
diff changeset
   181
Chengsong
parents:
diff changeset
   182
lemma r_in_rstar : "r1 \<leadsto> r2 \<Longrightarrow> r1 \<leadsto>* r2"
Chengsong
parents:
diff changeset
   183
  using rrewrites.intros(1) rrewrites.intros(2) by blast
Chengsong
parents:
diff changeset
   184
Chengsong
parents:
diff changeset
   185
lemma rs_in_rstar: 
Chengsong
parents:
diff changeset
   186
  shows "r1 s\<leadsto> r2 \<Longrightarrow> r1 s\<leadsto>* r2"
Chengsong
parents:
diff changeset
   187
  using rrewrites.intros(1) rrewrites.intros(2) by blast
Chengsong
parents:
diff changeset
   188
Chengsong
parents:
diff changeset
   189
Chengsong
parents:
diff changeset
   190
lemma rrewrites_trans[trans]: 
Chengsong
parents:
diff changeset
   191
  assumes a1: "r1 \<leadsto>* r2"  and a2: "r2 \<leadsto>* r3"
Chengsong
parents:
diff changeset
   192
  shows "r1 \<leadsto>* r3"
Chengsong
parents:
diff changeset
   193
  using a2 a1
Chengsong
parents:
diff changeset
   194
  apply(induct r2 r3 arbitrary: r1 rule: rrewrites.induct) 
Chengsong
parents:
diff changeset
   195
  apply(auto)
Chengsong
parents:
diff changeset
   196
  done
Chengsong
parents:
diff changeset
   197
Chengsong
parents:
diff changeset
   198
lemma srewrites_trans[trans]: 
Chengsong
parents:
diff changeset
   199
  assumes a1: "r1 s\<leadsto>* r2"  and a2: "r2 s\<leadsto>* r3"
Chengsong
parents:
diff changeset
   200
  shows "r1 s\<leadsto>* r3"
Chengsong
parents:
diff changeset
   201
  using a1 a2
Chengsong
parents:
diff changeset
   202
  apply(induct r1 r2 arbitrary: r3 rule: srewrites.induct) 
Chengsong
parents:
diff changeset
   203
   apply(auto)
Chengsong
parents:
diff changeset
   204
  done
Chengsong
parents:
diff changeset
   205
Chengsong
parents:
diff changeset
   206
lemma contextrewrites0: 
Chengsong
parents:
diff changeset
   207
  "rs1 s\<leadsto>* rs2 \<Longrightarrow> AALTs bs rs1 \<leadsto>* AALTs bs rs2"
Chengsong
parents:
diff changeset
   208
  apply(induct rs1 rs2 rule: srewrites.inducts)
Chengsong
parents:
diff changeset
   209
   apply simp
Chengsong
parents:
diff changeset
   210
  using bs10 r_in_rstar rrewrites_trans by blast
Chengsong
parents:
diff changeset
   211
Chengsong
parents:
diff changeset
   212
lemma contextrewrites1: 
Chengsong
parents:
diff changeset
   213
  "r \<leadsto>* r' \<Longrightarrow> AALTs bs (r # rs) \<leadsto>* AALTs bs (r' # rs)"
Chengsong
parents:
diff changeset
   214
  apply(induct r r' rule: rrewrites.induct)
Chengsong
parents:
diff changeset
   215
   apply simp
Chengsong
parents:
diff changeset
   216
  using bs10 ss3 by blast
Chengsong
parents:
diff changeset
   217
Chengsong
parents:
diff changeset
   218
lemma srewrite1: 
Chengsong
parents:
diff changeset
   219
  shows "rs1 s\<leadsto> rs2 \<Longrightarrow> (rs @ rs1) s\<leadsto> (rs @ rs2)"
Chengsong
parents:
diff changeset
   220
  apply(induct rs)
Chengsong
parents:
diff changeset
   221
   apply(auto)
Chengsong
parents:
diff changeset
   222
  using ss2 by auto
Chengsong
parents:
diff changeset
   223
Chengsong
parents:
diff changeset
   224
lemma srewrites1: 
Chengsong
parents:
diff changeset
   225
  shows "rs1 s\<leadsto>* rs2 \<Longrightarrow> (rs @ rs1) s\<leadsto>* (rs @ rs2)"
Chengsong
parents:
diff changeset
   226
  apply(induct rs1 rs2 rule: srewrites.induct)
Chengsong
parents:
diff changeset
   227
   apply(auto)
Chengsong
parents:
diff changeset
   228
  using srewrite1 by blast
Chengsong
parents:
diff changeset
   229
551
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   230
lemma srewrites_prepend:
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   231
  shows "rs1 s\<leadsto>* rs2 \<Longrightarrow> (r # rs1) s\<leadsto>* (r # rs2)"
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   232
  by (metis append_Cons append_Nil srewrites1)  
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   233
544
Chengsong
parents:
diff changeset
   234
lemma srewrite2: 
Chengsong
parents:
diff changeset
   235
  shows  "r1 \<leadsto> r2 \<Longrightarrow> True"
Chengsong
parents:
diff changeset
   236
  and "rs1 s\<leadsto> rs2 \<Longrightarrow> (rs1 @ rs) s\<leadsto>* (rs2 @ rs)"
551
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   237
  apply(induct arbitrary: rs rule: rrewrite_srewrite.inducts)
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   238
                apply simp+
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   239
  using srewrites_prepend apply force
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   240
      apply (simp add: rs_in_rstar ss3)
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   241
  using ss4 apply force 
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   242
  using rs_in_rstar ss5 apply auto[1]
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   243
  using rs_in_rstar ss6 apply auto[1]
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   244
  using rs_in_rstar ss7 by force
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   245
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   246
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   247
544
Chengsong
parents:
diff changeset
   248
Chengsong
parents:
diff changeset
   249
lemma srewrites3: 
Chengsong
parents:
diff changeset
   250
  shows "rs1 s\<leadsto>* rs2 \<Longrightarrow> (rs1 @ rs) s\<leadsto>* (rs2 @ rs)"
Chengsong
parents:
diff changeset
   251
  apply(induct rs1 rs2 arbitrary: rs rule: srewrites.induct)
Chengsong
parents:
diff changeset
   252
   apply(auto)
Chengsong
parents:
diff changeset
   253
  by (meson srewrite2(2) srewrites_trans)
Chengsong
parents:
diff changeset
   254
Chengsong
parents:
diff changeset
   255
Chengsong
parents:
diff changeset
   256
lemma srewrites6:
Chengsong
parents:
diff changeset
   257
  assumes "r1 \<leadsto>* r2" 
Chengsong
parents:
diff changeset
   258
  shows "[r1] s\<leadsto>* [r2]"
Chengsong
parents:
diff changeset
   259
  using assms
Chengsong
parents:
diff changeset
   260
  apply(induct r1 r2 rule: rrewrites.induct)
Chengsong
parents:
diff changeset
   261
   apply(auto)
Chengsong
parents:
diff changeset
   262
  by (meson srewrites.simps srewrites_trans ss3)
Chengsong
parents:
diff changeset
   263
Chengsong
parents:
diff changeset
   264
lemma srewrites7:
Chengsong
parents:
diff changeset
   265
  assumes "rs3 s\<leadsto>* rs4" "r1 \<leadsto>* r2" 
Chengsong
parents:
diff changeset
   266
  shows "(r1 # rs3) s\<leadsto>* (r2 # rs4)"
Chengsong
parents:
diff changeset
   267
  using assms
Chengsong
parents:
diff changeset
   268
  by (smt (verit, best) append_Cons append_Nil srewrites1 srewrites3 srewrites6 srewrites_trans)
Chengsong
parents:
diff changeset
   269
552
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   270
(*harmless sorry*)
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   271
lemma existing_preimage :
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   272
  shows "f a \<in> f ` set rs1 \<Longrightarrow> \<exists>rs1a rs1b x. rs1 = rs1a @ [x] @ rs1b \<and> f x = f a "
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   273
  apply(induct rs1)
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   274
   apply simp
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   275
  apply(case_tac "f a = f aa")
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   276
  
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   277
  sorry
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   278
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   279
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   280
lemma deletes_dB:
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   281
  shows " \<lbrakk>erase a \<in> erase ` set rs1\<rbrakk> \<Longrightarrow> (rs1 @ a # rs2) s\<leadsto>* (rs1 @ rs2)"
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   282
  apply(subgoal_tac "\<exists>rs1a rs1b x. rs1 = rs1a @ [x] @ rs1b \<and> erase x = erase a")
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   283
  prefer 2
553
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   284
   apply (meson existing_preimage)
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   285
  apply(erule exE)+
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   286
  apply simp
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   287
  apply(subgoal_tac "(rs1a @ [x] @ rs1b @ [a] @ rs2) s\<leadsto> (rs1a @ [x] @ rs1b @ rs2)")  
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   288
  apply (simp add: rs_in_rstar)
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   289
  apply(subgoal_tac "L (erase a) \<subseteq> L (erase x)")
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   290
  using ss6 apply presburger
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   291
  by simp
552
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   292
558
Chengsong
parents: 557
diff changeset
   293
Chengsong
parents: 557
diff changeset
   294
553
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   295
lemma ss6_realistic:
558
Chengsong
parents: 557
diff changeset
   296
  shows "(rs1 @ rs2) s\<leadsto>* (rs1 @ dB6 rs2 (tset rs1))"
557
812e5d112f49 more changes
Chengsong
parents: 553
diff changeset
   297
  apply(induct rs2 arbitrary: rs1)
812e5d112f49 more changes
Chengsong
parents: 553
diff changeset
   298
   apply simp
558
Chengsong
parents: 557
diff changeset
   299
  apply(rename_tac cc' cc)
Chengsong
parents: 557
diff changeset
   300
  apply(subgoal_tac "(cc @ a # cc') s\<leadsto>* (cc @ a # dB6 cc' (tset (cc @ [a])))")
Chengsong
parents: 557
diff changeset
   301
   prefer 2
Chengsong
parents: 557
diff changeset
   302
  apply (metis append.assoc append.left_neutral append_Cons)
564
Chengsong
parents: 558
diff changeset
   303
  apply(subgoal_tac "(cc @ a # dB6 cc' (tset (cc @ [a]))) s\<leadsto>* (cc @ (p6 cc a) # dB6 cc' (tset (cc @ [a])))") 
552
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   304
  sorry
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   305
564
Chengsong
parents: 558
diff changeset
   306
(*
544
Chengsong
parents:
diff changeset
   307
lemma ss6_stronger_aux:
551
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   308
  shows "(rs1 @ rs2) s\<leadsto>* (rs1 @ dB6 rs2 (set (map erase rs1)))"
544
Chengsong
parents:
diff changeset
   309
  apply(induct rs2 arbitrary: rs1)
552
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   310
   apply simp
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   311
  apply(case_tac "erase a \<in> erase ` set rs1")
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   312
   apply(simp)
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   313
   apply(drule_tac x = "rs1" in meta_spec)
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   314
  apply(subgoal_tac "(rs1 @ a # rs2) s\<leadsto>* (rs1 @ rs2)")
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   315
  using srewrites_trans apply blast
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   316
  using deletes_dB apply presburger
553
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   317
  apply(case_tac "set (turnIntoTerms (erase a)) \<subseteq> erase ` set rs1")
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   318
  
0f00d440f484 more changes
Chengsong
parents: 552
diff changeset
   319
  apply simp
552
3ea82d8f0aa4 start of day
Chengsong
parents: 551
diff changeset
   320
544
Chengsong
parents:
diff changeset
   321
  apply(auto)
Chengsong
parents:
diff changeset
   322
  prefer 2
Chengsong
parents:
diff changeset
   323
  apply(drule_tac x="rs1 @ [a]" in meta_spec)
Chengsong
parents:
diff changeset
   324
  apply(simp)
551
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   325
  sorry
564
Chengsong
parents: 558
diff changeset
   326
*)
551
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   327
544
Chengsong
parents:
diff changeset
   328
Chengsong
parents:
diff changeset
   329
lemma ss6_stronger:
551
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   330
  shows "rs1 s\<leadsto>* dB6 rs1 {}"
564
Chengsong
parents: 558
diff changeset
   331
  using ss6
Chengsong
parents: 558
diff changeset
   332
  by (metis append_Nil concat.simps(1) list.set(1) list.simps(8) ss6_realistic tset.simps)
Chengsong
parents: 558
diff changeset
   333
Chengsong
parents: 558
diff changeset
   334
 
Chengsong
parents: 558
diff changeset
   335
lemma tint_fuse:
Chengsong
parents: 558
diff changeset
   336
  shows "tint (erase a) = tint (erase (fuse bs a))"
Chengsong
parents: 558
diff changeset
   337
  by (simp add: erase_fuse)
Chengsong
parents: 558
diff changeset
   338
Chengsong
parents: 558
diff changeset
   339
lemma tint_fuse2:
Chengsong
parents: 558
diff changeset
   340
  shows " set (tint (erase a)) =
Chengsong
parents: 558
diff changeset
   341
     set (tint (erase (fuse bs a)))"
Chengsong
parents: 558
diff changeset
   342
  using tint_fuse by auto
Chengsong
parents: 558
diff changeset
   343
Chengsong
parents: 558
diff changeset
   344
lemma fused_bits_at_head:
Chengsong
parents: 558
diff changeset
   345
  shows "fuse bs a = ASEQ bs1 a1 a2 \<Longrightarrow> \<exists>bs2. bs1 = bs @ bs2"
Chengsong
parents: 558
diff changeset
   346
  
Chengsong
parents: 558
diff changeset
   347
(* by (smt (verit) arexp.distinct(13) arexp.distinct(19) arexp.distinct(25) arexp.distinct(27) arexp.distinct(5) arexp.inject(3) fuse.elims)
Chengsong
parents: 558
diff changeset
   348
harmless sorry
Chengsong
parents: 558
diff changeset
   349
*)
Chengsong
parents: 558
diff changeset
   350
551
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   351
  sorry
544
Chengsong
parents:
diff changeset
   352
564
Chengsong
parents: 558
diff changeset
   353
thm seqFold.simps
Chengsong
parents: 558
diff changeset
   354
Chengsong
parents: 558
diff changeset
   355
lemma seqFold_concats:
Chengsong
parents: 558
diff changeset
   356
  shows "r \<noteq> ONE \<Longrightarrow> seqFold r [r1] = SEQ r r1"
Chengsong
parents: 558
diff changeset
   357
  apply(case_tac r)
Chengsong
parents: 558
diff changeset
   358
       apply simp+
Chengsong
parents: 558
diff changeset
   359
done
Chengsong
parents: 558
diff changeset
   360
Chengsong
parents: 558
diff changeset
   361
568
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   362
lemma tint_seqFold_eq: shows"set (tint (seqFold (erase x42) [erase x43])) = 
564
Chengsong
parents: 558
diff changeset
   363
           set (tint (SEQ (erase x42) (erase x43)))"
Chengsong
parents: 558
diff changeset
   364
  apply(case_tac "erase x42 = ONE")
Chengsong
parents: 558
diff changeset
   365
   apply simp
568
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   366
  using seqFold_concats
564
Chengsong
parents: 558
diff changeset
   367
  apply simp
568
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   368
  done
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   369
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   370
fun top :: "arexp \<Rightarrow> bit list" where
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   371
  "top AZERO = []"
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   372
| "top (AONE bs) = bs"
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   373
| "top (ASEQ bs r1 r2) = bs"
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   374
| "top (ACHAR v va) = v"
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   375
| "top (AALTs v va) = v"
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   376
| "top (ASTAR v va) = v "
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   377
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   378
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   379
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   380
lemma top_bitcodes_preserved_p6:
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   381
  shows "\<lbrakk> r = ASEQ bs a1 a2 \<rbrakk> \<Longrightarrow> p6 as r = AZERO \<or> (\<exists>bs3. top (p6 as r) = bs @ bs3)"
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   382
  apply(induct r arbitrary: as)
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   383
       apply simp
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   384
  apply simp
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   385
     apply simp
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   386
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   387
  apply(case_tac "a1")
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   388
         apply simp
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   389
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   390
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   391
  sorry
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   392
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   393
564
Chengsong
parents: 558
diff changeset
   394
Chengsong
parents: 558
diff changeset
   395
lemma prune6_preserves_fuse:
Chengsong
parents: 558
diff changeset
   396
  shows "fuse bs (p6 as a) = p6 as (fuse bs a)"
Chengsong
parents: 558
diff changeset
   397
  using tint_fuse2
Chengsong
parents: 558
diff changeset
   398
  apply simp
Chengsong
parents: 558
diff changeset
   399
  apply(case_tac a)
Chengsong
parents: 558
diff changeset
   400
       apply simp
Chengsong
parents: 558
diff changeset
   401
  apply simp
Chengsong
parents: 558
diff changeset
   402
     apply simp
Chengsong
parents: 558
diff changeset
   403
Chengsong
parents: 558
diff changeset
   404
  using fused_bits_at_head
Chengsong
parents: 558
diff changeset
   405
Chengsong
parents: 558
diff changeset
   406
    apply simp
568
7a579f5533f8 more chapter2 modifications
Chengsong
parents: 564
diff changeset
   407
  using tint_seqFold_eq
564
Chengsong
parents: 558
diff changeset
   408
  apply simp
Chengsong
parents: 558
diff changeset
   409
Chengsong
parents: 558
diff changeset
   410
  sorry
Chengsong
parents: 558
diff changeset
   411
Chengsong
parents: 558
diff changeset
   412
corollary prune6_preserves_fuse_srewrite:
Chengsong
parents: 558
diff changeset
   413
  shows "(as @ map (fuse bs) [a] @ as2) s\<leadsto>* (as @ map (fuse bs) [p6 as a] @ as2)"
Chengsong
parents: 558
diff changeset
   414
  apply(subgoal_tac "map (fuse bs) [a] = [fuse bs a]")
Chengsong
parents: 558
diff changeset
   415
  apply(subgoal_tac "(as @ [fuse bs a] @ as2) s\<leadsto>* (as @ [ (p6 as (fuse bs a))] @ as2)")
Chengsong
parents: 558
diff changeset
   416
  using prune6_preserves_fuse apply auto[1]
Chengsong
parents: 558
diff changeset
   417
  using rs_in_rstar ss7 apply presburger
Chengsong
parents: 558
diff changeset
   418
  by simp
Chengsong
parents: 558
diff changeset
   419
Chengsong
parents: 558
diff changeset
   420
lemma prune6_invariant_fuse:
Chengsong
parents: 558
diff changeset
   421
  shows "p6 as a = p6 (map (fuse bs) as) a"
Chengsong
parents: 558
diff changeset
   422
  apply simp
Chengsong
parents: 558
diff changeset
   423
  using erase_fuse by presburger
Chengsong
parents: 558
diff changeset
   424
Chengsong
parents: 558
diff changeset
   425
Chengsong
parents: 558
diff changeset
   426
lemma p6pfs_cor:
Chengsong
parents: 558
diff changeset
   427
  shows "(map (fuse bs) as @ map (fuse bs) [a] @ map (fuse bs) as1) s\<leadsto>* (map (fuse bs) as @ map (fuse bs) [p6 as a] @ map (fuse bs) as1)"
Chengsong
parents: 558
diff changeset
   428
  by (metis prune6_invariant_fuse prune6_preserves_fuse_srewrite)
544
Chengsong
parents:
diff changeset
   429
Chengsong
parents:
diff changeset
   430
lemma rewrite_preserves_fuse: 
Chengsong
parents:
diff changeset
   431
  shows "r2 \<leadsto> r3 \<Longrightarrow> fuse bs r2 \<leadsto> fuse bs r3"
Chengsong
parents:
diff changeset
   432
  and   "rs2 s\<leadsto> rs3 \<Longrightarrow> map (fuse bs) rs2 s\<leadsto>* map (fuse bs) rs3"
Chengsong
parents:
diff changeset
   433
proof(induct rule: rrewrite_srewrite.inducts)
Chengsong
parents:
diff changeset
   434
  case (bs3 bs1 bs2 r)
Chengsong
parents:
diff changeset
   435
  then show ?case
Chengsong
parents:
diff changeset
   436
    by (metis fuse.simps(5) fuse_append rrewrite_srewrite.bs3) 
Chengsong
parents:
diff changeset
   437
next
Chengsong
parents:
diff changeset
   438
  case (bs7 bs r)
Chengsong
parents:
diff changeset
   439
  then show ?case
Chengsong
parents:
diff changeset
   440
    by (metis fuse.simps(4) fuse_append rrewrite_srewrite.bs7) 
Chengsong
parents:
diff changeset
   441
next
Chengsong
parents:
diff changeset
   442
  case (ss2 rs1 rs2 r)
Chengsong
parents:
diff changeset
   443
  then show ?case
Chengsong
parents:
diff changeset
   444
    using srewrites7 by force 
Chengsong
parents:
diff changeset
   445
next
Chengsong
parents:
diff changeset
   446
  case (ss3 r1 r2 rs)
Chengsong
parents:
diff changeset
   447
  then show ?case by (simp add: r_in_rstar srewrites7)
Chengsong
parents:
diff changeset
   448
next
Chengsong
parents:
diff changeset
   449
  case (ss5 bs1 rs1 rsb)
Chengsong
parents:
diff changeset
   450
  then show ?case 
Chengsong
parents:
diff changeset
   451
    apply(simp)
Chengsong
parents:
diff changeset
   452
    by (metis (mono_tags, lifting) comp_def fuse_append map_eq_conv rrewrite_srewrite.ss5 srewrites.simps)
Chengsong
parents:
diff changeset
   453
next
Chengsong
parents:
diff changeset
   454
  case (ss6 a1 a2 rsa rsb rsc)
Chengsong
parents:
diff changeset
   455
  then show ?case 
Chengsong
parents:
diff changeset
   456
    apply(simp only: map_append)
Chengsong
parents:
diff changeset
   457
    by (smt (verit, best) erase_fuse list.simps(8) list.simps(9) rrewrite_srewrite.ss6 srewrites.simps)
564
Chengsong
parents: 558
diff changeset
   458
next
Chengsong
parents: 558
diff changeset
   459
  case (ss7 as a as1)
Chengsong
parents: 558
diff changeset
   460
  then show ?case
Chengsong
parents: 558
diff changeset
   461
    apply(simp only: map_append)
Chengsong
parents: 558
diff changeset
   462
    using p6pfs_cor by presburger
544
Chengsong
parents:
diff changeset
   463
qed (auto intro: rrewrite_srewrite.intros)
564
Chengsong
parents: 558
diff changeset
   464
  
544
Chengsong
parents:
diff changeset
   465
Chengsong
parents:
diff changeset
   466
Chengsong
parents:
diff changeset
   467
lemma rewrites_fuse:  
Chengsong
parents:
diff changeset
   468
  assumes "r1 \<leadsto>* r2"
Chengsong
parents:
diff changeset
   469
  shows "fuse bs r1 \<leadsto>* fuse bs r2"
Chengsong
parents:
diff changeset
   470
using assms
Chengsong
parents:
diff changeset
   471
apply(induction r1 r2 arbitrary: bs rule: rrewrites.induct)
Chengsong
parents:
diff changeset
   472
apply(auto intro: rewrite_preserves_fuse rrewrites_trans)
Chengsong
parents:
diff changeset
   473
done
Chengsong
parents:
diff changeset
   474
Chengsong
parents:
diff changeset
   475
Chengsong
parents:
diff changeset
   476
lemma star_seq:  
Chengsong
parents:
diff changeset
   477
  assumes "r1 \<leadsto>* r2"
Chengsong
parents:
diff changeset
   478
  shows "ASEQ bs r1 r3 \<leadsto>* ASEQ bs r2 r3"
Chengsong
parents:
diff changeset
   479
using assms
Chengsong
parents:
diff changeset
   480
apply(induct r1 r2 arbitrary: r3 rule: rrewrites.induct)
Chengsong
parents:
diff changeset
   481
apply(auto intro: rrewrite_srewrite.intros)
Chengsong
parents:
diff changeset
   482
done
Chengsong
parents:
diff changeset
   483
Chengsong
parents:
diff changeset
   484
lemma star_seq2:  
Chengsong
parents:
diff changeset
   485
  assumes "r3 \<leadsto>* r4"
Chengsong
parents:
diff changeset
   486
  shows "ASEQ bs r1 r3 \<leadsto>* ASEQ bs r1 r4"
Chengsong
parents:
diff changeset
   487
  using assms
Chengsong
parents:
diff changeset
   488
apply(induct r3 r4 arbitrary: r1 rule: rrewrites.induct)
Chengsong
parents:
diff changeset
   489
apply(auto intro: rrewrite_srewrite.intros)
Chengsong
parents:
diff changeset
   490
done
Chengsong
parents:
diff changeset
   491
Chengsong
parents:
diff changeset
   492
lemma continuous_rewrite: 
Chengsong
parents:
diff changeset
   493
  assumes "r1 \<leadsto>* AZERO"
Chengsong
parents:
diff changeset
   494
  shows "ASEQ bs1 r1 r2 \<leadsto>* AZERO"
Chengsong
parents:
diff changeset
   495
using assms bs1 star_seq by blast
Chengsong
parents:
diff changeset
   496
551
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   497
544
Chengsong
parents:
diff changeset
   498
Chengsong
parents:
diff changeset
   499
lemma bsimp_aalts_simpcases: 
551
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   500
  shows "AONE bs \<leadsto>* bsimpStrong6 (AONE bs)"  
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   501
  and   "AZERO \<leadsto>* bsimpStrong6 AZERO" 
1243a031966c modified some proofs of s~>*
Chengsong
parents: 550
diff changeset
   502
  and   "ACHAR bs c \<leadsto>* bsimpStrong6 (ACHAR bs c)"
544
Chengsong
parents:
diff changeset
   503
  by (simp_all)
Chengsong
parents:
diff changeset
   504
Chengsong
parents:
diff changeset
   505
lemma bsimp_AALTs_rewrites: 
Chengsong
parents:
diff changeset
   506
  shows "AALTs bs1 rs \<leadsto>* bsimp_AALTs bs1 rs"
Chengsong
parents:
diff changeset
   507
  by (smt (verit) bs6 bs7 bsimp_AALTs.elims rrewrites.simps)
Chengsong
parents:
diff changeset
   508
Chengsong
parents:
diff changeset
   509
lemma trivialbsimp_srewrites: 
Chengsong
parents:
diff changeset
   510
  "\<lbrakk>\<And>x. x \<in> set rs \<Longrightarrow> x \<leadsto>* f x \<rbrakk> \<Longrightarrow> rs s\<leadsto>* (map f rs)"
Chengsong
parents:
diff changeset
   511
  apply(induction rs)
Chengsong
parents:
diff changeset
   512
   apply simp
Chengsong
parents:
diff changeset
   513
  apply(simp)
Chengsong
parents:
diff changeset
   514
  using srewrites7 by auto
Chengsong
parents:
diff changeset
   515
Chengsong
parents:
diff changeset
   516
Chengsong
parents:
diff changeset
   517
Chengsong
parents:
diff changeset
   518
lemma fltsfrewrites: "rs s\<leadsto>* flts rs"
Chengsong
parents:
diff changeset
   519
  apply(induction rs rule: flts.induct)
Chengsong
parents:
diff changeset
   520
  apply(auto intro: rrewrite_srewrite.intros)
Chengsong
parents:
diff changeset
   521
  apply (meson srewrites.simps srewrites1 ss5)
Chengsong
parents:
diff changeset
   522
  using rs1 srewrites7 apply presburger
Chengsong
parents:
diff changeset
   523
  using srewrites7 apply force
Chengsong
parents:
diff changeset
   524
  apply (simp add: srewrites7)
Chengsong
parents:
diff changeset
   525
  by (simp add: srewrites7)
Chengsong
parents:
diff changeset
   526
Chengsong
parents:
diff changeset
   527
lemma bnullable0:
Chengsong
parents:
diff changeset
   528
shows "r1 \<leadsto> r2 \<Longrightarrow> bnullable r1 = bnullable r2" 
Chengsong
parents:
diff changeset
   529
  and "rs1 s\<leadsto> rs2 \<Longrightarrow> bnullables rs1 = bnullables rs2" 
564
Chengsong
parents: 558
diff changeset
   530
   apply(induct rule: rrewrite_srewrite.inducts)
Chengsong
parents: 558
diff changeset
   531
              apply simp
Chengsong
parents: 558
diff changeset
   532
             apply simp
Chengsong
parents: 558
diff changeset
   533
              apply (simp add: bnullable_fuse)
Chengsong
parents: 558
diff changeset
   534
  using bnullable.simps(5) apply presburger
Chengsong
parents: 558
diff changeset
   535
          apply simp
Chengsong
parents: 558
diff changeset
   536
         apply simp
Chengsong
parents: 558
diff changeset
   537
  sorry   
548
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   538
544
Chengsong
parents:
diff changeset
   539
Chengsong
parents:
diff changeset
   540
Chengsong
parents:
diff changeset
   541
lemma rewritesnullable: 
Chengsong
parents:
diff changeset
   542
  assumes "r1 \<leadsto>* r2" 
Chengsong
parents:
diff changeset
   543
  shows "bnullable r1 = bnullable r2"
Chengsong
parents:
diff changeset
   544
using assms 
Chengsong
parents:
diff changeset
   545
  apply(induction r1 r2 rule: rrewrites.induct)
Chengsong
parents:
diff changeset
   546
  apply simp
564
Chengsong
parents: 558
diff changeset
   547
  using bnullable0(1) by presburger
544
Chengsong
parents:
diff changeset
   548
Chengsong
parents:
diff changeset
   549
lemma rewrite_bmkeps_aux: 
Chengsong
parents:
diff changeset
   550
  shows "r1 \<leadsto> r2 \<Longrightarrow> (bnullable r1 \<and> bnullable r2 \<Longrightarrow> bmkeps r1 = bmkeps r2)"
Chengsong
parents:
diff changeset
   551
  and   "rs1 s\<leadsto> rs2 \<Longrightarrow> (bnullables rs1 \<and> bnullables rs2 \<Longrightarrow> bmkepss rs1 = bmkepss rs2)" 
Chengsong
parents:
diff changeset
   552
proof (induct rule: rrewrite_srewrite.inducts)
Chengsong
parents:
diff changeset
   553
  case (bs3 bs1 bs2 r)
Chengsong
parents:
diff changeset
   554
  then show ?case by (simp add: bmkeps_fuse) 
Chengsong
parents:
diff changeset
   555
next
Chengsong
parents:
diff changeset
   556
  case (bs7 bs r)
Chengsong
parents:
diff changeset
   557
  then show ?case by (simp add: bmkeps_fuse) 
Chengsong
parents:
diff changeset
   558
next
Chengsong
parents:
diff changeset
   559
  case (ss3 r1 r2 rs)
Chengsong
parents:
diff changeset
   560
  then show ?case
Chengsong
parents:
diff changeset
   561
    using bmkepss.simps bnullable0(1) by presburger
Chengsong
parents:
diff changeset
   562
next
Chengsong
parents:
diff changeset
   563
  case (ss5 bs1 rs1 rsb)
Chengsong
parents:
diff changeset
   564
  then show ?case
Chengsong
parents:
diff changeset
   565
    by (simp add: bmkepss1 bmkepss2 bmkepss_fuse bnullable_fuse)
Chengsong
parents:
diff changeset
   566
next
Chengsong
parents:
diff changeset
   567
  case (ss6 a1 a2 rsa rsb rsc)
Chengsong
parents:
diff changeset
   568
  then show ?case
Chengsong
parents:
diff changeset
   569
    by (smt (verit, best) Nil_is_append_conv bmkepss1 bmkepss2 bnullable_correctness in_set_conv_decomp list.distinct(1) list.set_intros(1) nullable_correctness set_ConsD subsetD)
Chengsong
parents:
diff changeset
   570
qed (auto)
Chengsong
parents:
diff changeset
   571
Chengsong
parents:
diff changeset
   572
lemma rewrites_bmkeps: 
Chengsong
parents:
diff changeset
   573
  assumes "r1 \<leadsto>* r2" "bnullable r1" 
Chengsong
parents:
diff changeset
   574
  shows "bmkeps r1 = bmkeps r2"
Chengsong
parents:
diff changeset
   575
  using assms
Chengsong
parents:
diff changeset
   576
proof(induction r1 r2 rule: rrewrites.induct)
Chengsong
parents:
diff changeset
   577
  case (rs1 r)
Chengsong
parents:
diff changeset
   578
  then show "bmkeps r = bmkeps r" by simp
Chengsong
parents:
diff changeset
   579
next
Chengsong
parents:
diff changeset
   580
  case (rs2 r1 r2 r3)
Chengsong
parents:
diff changeset
   581
  then have IH: "bmkeps r1 = bmkeps r2" by simp
Chengsong
parents:
diff changeset
   582
  have a1: "bnullable r1" by fact
Chengsong
parents:
diff changeset
   583
  have a2: "r1 \<leadsto>* r2" by fact
Chengsong
parents:
diff changeset
   584
  have a3: "r2 \<leadsto> r3" by fact
Chengsong
parents:
diff changeset
   585
  have a4: "bnullable r2" using a1 a2 by (simp add: rewritesnullable) 
Chengsong
parents:
diff changeset
   586
  then have "bmkeps r2 = bmkeps r3"
Chengsong
parents:
diff changeset
   587
    using a3 bnullable0(1) rewrite_bmkeps_aux(1) by blast 
Chengsong
parents:
diff changeset
   588
  then show "bmkeps r1 = bmkeps r3" using IH by simp
Chengsong
parents:
diff changeset
   589
qed
Chengsong
parents:
diff changeset
   590
Chengsong
parents:
diff changeset
   591
Chengsong
parents:
diff changeset
   592
Chengsong
parents:
diff changeset
   593
lemma to_zero_in_alt: 
Chengsong
parents:
diff changeset
   594
  shows "AALT bs (ASEQ [] AZERO r) r2 \<leadsto> AALT bs AZERO r2"
Chengsong
parents:
diff changeset
   595
  by (simp add: bs1 bs10 ss3)
Chengsong
parents:
diff changeset
   596
Chengsong
parents:
diff changeset
   597
Chengsong
parents:
diff changeset
   598
Chengsong
parents:
diff changeset
   599
lemma  bder_fuse_list: 
Chengsong
parents:
diff changeset
   600
  shows "map (bder c \<circ> fuse bs1) rs1 = map (fuse bs1 \<circ> bder c) rs1"
Chengsong
parents:
diff changeset
   601
  apply(induction rs1)
Chengsong
parents:
diff changeset
   602
  apply(simp_all add: bder_fuse)
Chengsong
parents:
diff changeset
   603
  done
Chengsong
parents:
diff changeset
   604
Chengsong
parents:
diff changeset
   605
lemma map1:
Chengsong
parents:
diff changeset
   606
  shows "(map f [a]) = [f a]"
Chengsong
parents:
diff changeset
   607
  by (simp)
Chengsong
parents:
diff changeset
   608
Chengsong
parents:
diff changeset
   609
lemma rewrite_preserves_bder: 
Chengsong
parents:
diff changeset
   610
  shows "r1 \<leadsto> r2 \<Longrightarrow> (bder c r1) \<leadsto>* (bder c r2)"
Chengsong
parents:
diff changeset
   611
  and   "rs1 s\<leadsto> rs2 \<Longrightarrow> map (bder c) rs1 s\<leadsto>* map (bder c) rs2"
Chengsong
parents:
diff changeset
   612
proof(induction rule: rrewrite_srewrite.inducts)
Chengsong
parents:
diff changeset
   613
  case (bs1 bs r2)
Chengsong
parents:
diff changeset
   614
  then show ?case
Chengsong
parents:
diff changeset
   615
    by (simp add: continuous_rewrite) 
Chengsong
parents:
diff changeset
   616
next
Chengsong
parents:
diff changeset
   617
  case (bs2 bs r1)
Chengsong
parents:
diff changeset
   618
  then show ?case 
Chengsong
parents:
diff changeset
   619
    apply(auto)
Chengsong
parents:
diff changeset
   620
    apply (meson bs6 contextrewrites0 rrewrite_srewrite.bs2 rs2 ss3 ss4 sss1 sss2)
Chengsong
parents:
diff changeset
   621
    by (simp add: r_in_rstar rrewrite_srewrite.bs2)
Chengsong
parents:
diff changeset
   622
next
Chengsong
parents:
diff changeset
   623
  case (bs3 bs1 bs2 r)
Chengsong
parents:
diff changeset
   624
  then show ?case 
Chengsong
parents:
diff changeset
   625
    apply(simp)
Chengsong
parents:
diff changeset
   626
    
Chengsong
parents:
diff changeset
   627
    by (metis (no_types, lifting) bder_fuse bs10 bs7 fuse_append rrewrites.simps ss4 to_zero_in_alt)
Chengsong
parents:
diff changeset
   628
next
Chengsong
parents:
diff changeset
   629
  case (bs4 r1 r2 bs r3)
Chengsong
parents:
diff changeset
   630
  have as: "r1 \<leadsto> r2" by fact
Chengsong
parents:
diff changeset
   631
  have IH: "bder c r1 \<leadsto>* bder c r2" by fact
Chengsong
parents:
diff changeset
   632
  from as IH show "bder c (ASEQ bs r1 r3) \<leadsto>* bder c (ASEQ bs r2 r3)"
Chengsong
parents:
diff changeset
   633
    by (metis bder.simps(5) bnullable0(1) contextrewrites1 rewrite_bmkeps_aux(1) star_seq)
Chengsong
parents:
diff changeset
   634
next
Chengsong
parents:
diff changeset
   635
  case (bs5 r3 r4 bs r1)
Chengsong
parents:
diff changeset
   636
  have as: "r3 \<leadsto> r4" by fact 
Chengsong
parents:
diff changeset
   637
  have IH: "bder c r3 \<leadsto>* bder c r4" by fact 
Chengsong
parents:
diff changeset
   638
  from as IH show "bder c (ASEQ bs r1 r3) \<leadsto>* bder c (ASEQ bs r1 r4)"
Chengsong
parents:
diff changeset
   639
    apply(simp)
Chengsong
parents:
diff changeset
   640
    apply(auto)
Chengsong
parents:
diff changeset
   641
    using contextrewrites0 r_in_rstar rewrites_fuse srewrites6 srewrites7 star_seq2 apply presburger
Chengsong
parents:
diff changeset
   642
    using star_seq2 by blast
Chengsong
parents:
diff changeset
   643
next
Chengsong
parents:
diff changeset
   644
  case (bs6 bs)
Chengsong
parents:
diff changeset
   645
  then show ?case
Chengsong
parents:
diff changeset
   646
    using rrewrite_srewrite.bs6 by force 
Chengsong
parents:
diff changeset
   647
next
Chengsong
parents:
diff changeset
   648
  case (bs7 bs r)
Chengsong
parents:
diff changeset
   649
  then show ?case
Chengsong
parents:
diff changeset
   650
    by (simp add: bder_fuse r_in_rstar rrewrite_srewrite.bs7) 
Chengsong
parents:
diff changeset
   651
next
Chengsong
parents:
diff changeset
   652
  case (bs10 rs1 rs2 bs)
Chengsong
parents:
diff changeset
   653
  then show ?case
Chengsong
parents:
diff changeset
   654
    using contextrewrites0 by force    
Chengsong
parents:
diff changeset
   655
next
Chengsong
parents:
diff changeset
   656
  case ss1
Chengsong
parents:
diff changeset
   657
  then show ?case by simp
Chengsong
parents:
diff changeset
   658
next
Chengsong
parents:
diff changeset
   659
  case (ss2 rs1 rs2 r)
Chengsong
parents:
diff changeset
   660
  then show ?case
Chengsong
parents:
diff changeset
   661
    by (simp add: srewrites7) 
Chengsong
parents:
diff changeset
   662
next
Chengsong
parents:
diff changeset
   663
  case (ss3 r1 r2 rs)
Chengsong
parents:
diff changeset
   664
  then show ?case
Chengsong
parents:
diff changeset
   665
    by (simp add: srewrites7) 
Chengsong
parents:
diff changeset
   666
next
Chengsong
parents:
diff changeset
   667
  case (ss4 rs)
Chengsong
parents:
diff changeset
   668
  then show ?case
Chengsong
parents:
diff changeset
   669
    using rrewrite_srewrite.ss4 by fastforce 
Chengsong
parents:
diff changeset
   670
next
Chengsong
parents:
diff changeset
   671
  case (ss5 bs1 rs1 rsb)
Chengsong
parents:
diff changeset
   672
  then show ?case 
Chengsong
parents:
diff changeset
   673
    apply(simp)
Chengsong
parents:
diff changeset
   674
    using bder_fuse_list map_map rrewrite_srewrite.ss5 srewrites.simps by blast
Chengsong
parents:
diff changeset
   675
next
Chengsong
parents:
diff changeset
   676
  case (ss6 a1 a2 bs rsa rsb)
Chengsong
parents:
diff changeset
   677
  then show ?case
Chengsong
parents:
diff changeset
   678
    apply(simp only: map_append map1)
Chengsong
parents:
diff changeset
   679
    apply(rule srewrites_trans)
Chengsong
parents:
diff changeset
   680
    apply(rule rs_in_rstar)
Chengsong
parents:
diff changeset
   681
    apply(rule_tac rrewrite_srewrite.ss6)
Chengsong
parents:
diff changeset
   682
    using Der_def der_correctness apply auto[1]
Chengsong
parents:
diff changeset
   683
    by blast
Chengsong
parents:
diff changeset
   684
qed
Chengsong
parents:
diff changeset
   685
Chengsong
parents:
diff changeset
   686
lemma rewrites_preserves_bder: 
Chengsong
parents:
diff changeset
   687
  assumes "r1 \<leadsto>* r2"
Chengsong
parents:
diff changeset
   688
  shows "bder c r1 \<leadsto>* bder c r2"
Chengsong
parents:
diff changeset
   689
using assms  
Chengsong
parents:
diff changeset
   690
apply(induction r1 r2 rule: rrewrites.induct)
Chengsong
parents:
diff changeset
   691
apply(simp_all add: rewrite_preserves_bder rrewrites_trans)
548
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   692
  done
544
Chengsong
parents:
diff changeset
   693
Chengsong
parents:
diff changeset
   694
Chengsong
parents:
diff changeset
   695
548
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   696
lemma bders_simp_appendStrong :
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   697
  shows "bdersStrong6 r (s1 @ s2) = bdersStrong6 (bdersStrong6 r s1) s2"
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   698
  apply(induct s1 arbitrary: s2 rule: rev_induct)
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   699
   apply simp
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   700
  apply auto
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   701
  done
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   702
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   703
546
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   704
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   705
547
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   706
lemma rewrites_to_bsimpStrong: 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   707
  shows "r \<leadsto>* bsimpStrong6 r"
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   708
proof (induction r rule: bsimpStrong6.induct)
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   709
  case (1 bs1 r1 r2)
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   710
  have IH1: "r1 \<leadsto>* bsimpStrong6 r1" by fact
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   711
  have IH2: "r2 \<leadsto>* bsimpStrong6 r2" by fact
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   712
  { assume as: "bsimpStrong6 r1 = AZERO \<or> bsimpStrong6 r2 = AZERO"
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   713
    with IH1 IH2 have "r1 \<leadsto>* AZERO \<or> r2 \<leadsto>* AZERO" by auto
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   714
    then have "ASEQ bs1 r1 r2 \<leadsto>* AZERO"
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   715
      by (metis bs2 continuous_rewrite rrewrites.simps star_seq2)  
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   716
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimpStrong6 (ASEQ bs1 r1 r2)" using as by auto
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   717
  }
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   718
  moreover
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   719
  { assume "\<exists>bs. bsimpStrong6 r1 = AONE bs"
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   720
    then obtain bs where as: "bsimpStrong6 r1 = AONE bs" by blast
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   721
    with IH1 have "r1 \<leadsto>* AONE bs" by simp
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   722
    then have "ASEQ bs1 r1 r2 \<leadsto>* fuse (bs1 @ bs) r2" using bs3 star_seq by blast
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   723
    with IH2 have "ASEQ bs1 r1 r2 \<leadsto>* fuse (bs1 @ bs) (bsimpStrong6 r2)"
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   724
      using rewrites_fuse by (meson rrewrites_trans) 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   725
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimpStrong6 (ASEQ bs1 (AONE bs) r2)" by simp
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   726
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimpStrong6 (ASEQ bs1 r1 r2)" by (simp add: as) 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   727
  } 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   728
  moreover
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   729
  { assume as1: "bsimpStrong6 r1 \<noteq> AZERO" "bsimpStrong6 r2 \<noteq> AZERO" and as2: "(\<nexists>bs. bsimpStrong6 r1 = AONE bs)" 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   730
    then have "bsimp_ASEQ bs1 (bsimpStrong6 r1) (bsimpStrong6 r2) = ASEQ bs1 (bsimpStrong6 r1) (bsimpStrong6 r2)" 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   731
      by (simp add: bsimp_ASEQ1) 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   732
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp_ASEQ bs1 (bsimpStrong6 r1) (bsimpStrong6 r2)" using as1 as2 IH1 IH2
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   733
      by (metis rrewrites_trans star_seq star_seq2) 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   734
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimpStrong6 (ASEQ bs1 r1 r2)" by simp
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   735
  } 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   736
  ultimately show "ASEQ bs1 r1 r2 \<leadsto>* bsimpStrong6 (ASEQ bs1 r1 r2)" sorry
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   737
next
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   738
  case (2 bs1 rs)
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   739
  have IH: "\<And>x. x \<in> set rs \<Longrightarrow> x \<leadsto>* bsimpStrong6 x" by fact
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   740
  then have "rs s\<leadsto>* (map bsimpStrong6 rs)" by (simp add: trivialbsimp_srewrites)
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   741
  also have "... s\<leadsto>* flts (map bsimpStrong6 rs)" by (simp add: fltsfrewrites) 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   742
  also have "... s\<leadsto>* distinctWith (flts (map bsimpStrong6 rs)) eq1 {}" by (simp add: ss6_stronger)
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   743
  finally have "AALTs bs1 rs \<leadsto>* AALTs bs1 (distinctWith (flts (map bsimpStrong6 rs)) eq1 {})"
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   744
    using contextrewrites0 by auto
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   745
  also have "... \<leadsto>* bsimp_AALTs  bs1 (distinctWith (flts (map bsimpStrong6 rs)) eq1 {})"
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   746
    by (simp add: bsimp_AALTs_rewrites)     
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   747
  finally show "AALTs bs1 rs \<leadsto>* bsimpStrong6 (AALTs bs1 rs)" sorry
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   748
qed (simp_all)
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   749
548
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   750
547
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   751
546
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   752
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   753
lemma centralStrong:  
546
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   754
  shows "bders r s \<leadsto>* bdersStrong6 r s"
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   755
proof(induct s arbitrary: r rule: rev_induct)
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   756
  case Nil
546
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   757
  then show "bders r [] \<leadsto>* bdersStrong6 r []" by simp
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   758
next
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   759
  case (snoc x xs)
546
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   760
  have IH: "\<And>r. bders r xs \<leadsto>* bdersStrong6 r xs" by fact
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   761
  have "bders r (xs @ [x]) = bders (bders r xs) [x]" by (simp add: bders_append)
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   762
  also have "... \<leadsto>* bders (bdersStrong6 r xs) [x]" using IH
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   763
    by (simp add: rewrites_preserves_bder)
546
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   764
  also have "... \<leadsto>* bdersStrong6 (bdersStrong6 r xs) [x]" using IH
547
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   765
    by (simp add: rewrites_to_bsimpStrong)
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   766
  finally show "bders r (xs @ [x]) \<leadsto>* bdersStrong6 r (xs @ [x])" 
feae84f66472 before alternating rewriting relation
Chengsong
parents: 546
diff changeset
   767
    by (simp add: bders_simp_appendStrong)
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   768
qed
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   769
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   770
lemma mainStrong: 
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   771
  assumes "bnullable (bders r s)"
546
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   772
  shows "bmkeps (bders r s) = bmkeps (bdersStrong6 r s)"
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   773
proof -
548
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   774
  have "bders r s \<leadsto>* bdersStrong6 r s" 
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   775
    using centralStrong by force
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   776
  then 
546
6e97f4aa7cd0 beforeBig changes
Chengsong
parents: 545
diff changeset
   777
  show "bmkeps (bders r s) = bmkeps (bdersStrong6 r s)" 
548
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   778
    using assms rewrites_bmkeps by blast
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   779
qed
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   780
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   781
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   782
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   783
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   784
theorem blexerStrong_correct :
548
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   785
  shows "blexerStrong r s = blexer r s"
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   786
  unfolding blexerStrong_def blexer_def
e5db23c100ea before removing distinctWith
Chengsong
parents: 547
diff changeset
   787
  by (metis centralStrong mainStrong rewritesnullable)
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 544
diff changeset
   788
544
Chengsong
parents:
diff changeset
   789
Chengsong
parents:
diff changeset
   790
Chengsong
parents:
diff changeset
   791
end