| 444 |      1 | theory ClosedForms imports
 | 
|  |      2 | "BasicIdentities"
 | 
| 443 |      3 | begin
 | 
|  |      4 | 
 | 
| 453 |      5 | 
 | 
| 465 |      6 | lemma idem_after_simp1:
 | 
|  |      7 |   shows "rsimp_ALTs (rdistinct (rflts [rsimp aa]) {}) = rsimp aa"
 | 
|  |      8 |   apply(case_tac "rsimp aa")
 | 
|  |      9 |   apply simp+
 | 
|  |     10 |   apply (metis no_alt_short_list_after_simp no_further_dB_after_simp)
 | 
|  |     11 |   by simp
 | 
| 456 |     12 | 
 | 
|  |     13 | 
 | 
|  |     14 | lemma distinct_removes_last:
 | 
| 465 |     15 |   shows "\<lbrakk>a \<in> set as\<rbrakk>
 | 
| 456 |     16 |     \<Longrightarrow> rdistinct as rset = rdistinct (as @ [a]) rset"
 | 
|  |     17 | and "rdistinct (ab # as @ [ab]) rset1 = rdistinct (ab # as) rset1"
 | 
|  |     18 |   apply(induct as arbitrary: rset ab rset1 a)
 | 
|  |     19 |      apply simp
 | 
|  |     20 |     apply simp
 | 
|  |     21 |   apply(case_tac "aa \<in> rset")
 | 
|  |     22 |    apply(case_tac "a = aa")
 | 
|  |     23 |   apply (metis append_Cons)
 | 
|  |     24 |     apply simp
 | 
|  |     25 |    apply(case_tac "a \<in> set as")
 | 
|  |     26 |   apply (metis append_Cons rdistinct.simps(2) set_ConsD)
 | 
|  |     27 |    apply(case_tac "a = aa")
 | 
|  |     28 |     prefer 2
 | 
|  |     29 |     apply simp
 | 
|  |     30 |    apply (metis append_Cons)
 | 
|  |     31 |   apply(case_tac "ab \<in> rset1")
 | 
|  |     32 |   prefer 2
 | 
|  |     33 |    apply(subgoal_tac "rdistinct (ab # (a # as) @ [ab]) rset1 = 
 | 
|  |     34 |                ab # (rdistinct ((a # as) @ [ab]) (insert ab rset1))")
 | 
|  |     35 |   prefer 2
 | 
|  |     36 |   apply force
 | 
|  |     37 |   apply(simp only:)
 | 
|  |     38 |      apply(subgoal_tac "rdistinct (ab # a # as) rset1 = ab # (rdistinct (a # as) (insert ab rset1))")
 | 
|  |     39 |     apply(simp only:)
 | 
|  |     40 |     apply(subgoal_tac "rdistinct ((a # as) @ [ab]) (insert ab rset1) = rdistinct (a # as) (insert ab rset1)")
 | 
|  |     41 |      apply blast
 | 
|  |     42 |     apply(case_tac "a \<in> insert ab rset1")
 | 
|  |     43 |      apply simp
 | 
|  |     44 |      apply (metis insertI1)
 | 
|  |     45 |     apply simp
 | 
|  |     46 |     apply (meson insertI1)
 | 
|  |     47 |    apply simp
 | 
|  |     48 |   apply(subgoal_tac "rdistinct ((a # as) @ [ab]) rset1 = rdistinct (a # as) rset1")
 | 
|  |     49 |    apply simp
 | 
|  |     50 |   by (metis append_Cons insert_iff insert_is_Un rdistinct.simps(2))
 | 
|  |     51 | 
 | 
|  |     52 | 
 | 
| 465 |     53 | lemma distinct_removes_middle:
 | 
|  |     54 |   shows  "\<lbrakk>a \<in> set as\<rbrakk>
 | 
|  |     55 |     \<Longrightarrow> rdistinct (as @ as2) rset = rdistinct (as @ [a] @ as2) rset"
 | 
|  |     56 | and "rdistinct (ab # as @ [ab] @ as3) rset1 = rdistinct (ab # as @ as3) rset1"
 | 
|  |     57 |    apply(induct as arbitrary: rset rset1 ab as2 as3 a)
 | 
|  |     58 |      apply simp
 | 
|  |     59 |     apply simp
 | 
|  |     60 |    apply(case_tac "a \<in> rset")
 | 
|  |     61 |     apply simp
 | 
|  |     62 |     apply metis
 | 
| 453 |     63 |    apply simp
 | 
| 465 |     64 |    apply (metis insertI1)
 | 
|  |     65 |   apply(case_tac "a = ab")
 | 
| 453 |     66 |    apply simp
 | 
| 465 |     67 |    apply(case_tac "ab \<in> rset")
 | 
| 453 |     68 |     apply simp
 | 
| 465 |     69 |     apply presburger
 | 
|  |     70 |    apply (meson insertI1)
 | 
|  |     71 |   apply(case_tac "a \<in> rset")
 | 
|  |     72 |   apply (metis (no_types, opaque_lifting) Un_insert_left append_Cons insert_iff rdistinct.simps(2) sup_bot_left)
 | 
|  |     73 |   apply(case_tac "ab \<in> rset")
 | 
|  |     74 |   apply simp
 | 
|  |     75 |    apply (meson insert_iff)
 | 
|  |     76 |   apply simp
 | 
|  |     77 |   by (metis insertI1)
 | 
| 453 |     78 | 
 | 
|  |     79 | 
 | 
| 465 |     80 | lemma distinct_removes_middle3:
 | 
|  |     81 |   shows  "\<lbrakk>a \<in> set as\<rbrakk>
 | 
|  |     82 |     \<Longrightarrow> rdistinct (as @ a #as2) rset = rdistinct (as @ as2) rset"
 | 
|  |     83 |   using distinct_removes_middle(1) by fastforce
 | 
|  |     84 | 
 | 
|  |     85 | lemma distinct_removes_last2:
 | 
|  |     86 |   shows "\<lbrakk>a \<in> set as\<rbrakk>
 | 
|  |     87 |     \<Longrightarrow> rdistinct as rset = rdistinct (as @ [a]) rset"
 | 
|  |     88 |   by (simp add: distinct_removes_last(1))
 | 
|  |     89 | 
 | 
|  |     90 | lemma distinct_removes_middle2:
 | 
|  |     91 |   shows "a \<in> set as \<Longrightarrow> rdistinct (as @ [a] @ rs) {} = rdistinct (as @ rs) {}"
 | 
|  |     92 |   by (metis distinct_removes_middle(1))
 | 
|  |     93 | 
 | 
|  |     94 | lemma distinct_removes_list:
 | 
|  |     95 |   shows "\<lbrakk>a \<in> set as; \<forall>r \<in> set rs. r \<in> set as\<rbrakk> \<Longrightarrow> rdistinct (as @ rs) {} = rdistinct as {}"
 | 
|  |     96 |   apply(induct rs)
 | 
|  |     97 |    apply simp+
 | 
|  |     98 |   apply(subgoal_tac "rdistinct (as @ aa # rs) {} = rdistinct (as @ rs) {}")
 | 
|  |     99 |    prefer 2
 | 
|  |    100 |   apply (metis append_Cons append_Nil distinct_removes_middle(1))
 | 
|  |    101 |   by presburger
 | 
| 453 |    102 | 
 | 
| 451 |    103 | 
 | 
|  |    104 | 
 | 
|  |    105 | lemma simp_rdistinct_f: shows 
 | 
| 465 |    106 | "f ` rset = frset \<Longrightarrow> rsimp (rsimp_ALTs (map f (rdistinct rs rset))) = 
 | 
|  |    107 |                       rsimp (rsimp_ALTs (rdistinct (map f rs) frset))  "
 | 
| 451 |    108 |   apply(induct rs arbitrary: rset)
 | 
|  |    109 |    apply simp
 | 
|  |    110 |    apply(case_tac "a \<in> rset")
 | 
|  |    111 |   apply(case_tac " f a \<in> frset")
 | 
|  |    112 |    apply simp
 | 
|  |    113 |    apply blast
 | 
|  |    114 |   apply(subgoal_tac "f a \<notin> frset")
 | 
|  |    115 |    apply(simp)
 | 
|  |    116 |    apply(subgoal_tac "f ` (insert a rset) = insert (f a) frset")
 | 
|  |    117 |   prefer 2
 | 
|  |    118 |   apply (meson image_insert)
 | 
|  |    119 |   
 | 
| 453 |    120 |   oops
 | 
| 451 |    121 | 
 | 
| 453 |    122 | lemma spawn_simp_rsimpalts:
 | 
|  |    123 |   shows "rsimp (rsimp_ALTs rs) = rsimp (rsimp_ALTs (map rsimp rs))"
 | 
|  |    124 |   apply(cases rs)
 | 
|  |    125 |    apply simp
 | 
|  |    126 |   apply(case_tac list)
 | 
|  |    127 |    apply simp
 | 
|  |    128 |    apply(subst rsimp_idem[symmetric])
 | 
|  |    129 |    apply simp
 | 
|  |    130 |   apply(subgoal_tac "rsimp_ALTs rs = RALTS rs")
 | 
|  |    131 |    apply(simp only:)
 | 
|  |    132 |    apply(subgoal_tac "rsimp_ALTs (map rsimp rs) = RALTS (map rsimp rs)")
 | 
|  |    133 |     apply(simp only:)
 | 
|  |    134 |   prefer 2
 | 
|  |    135 |   apply simp
 | 
|  |    136 |    prefer 2
 | 
|  |    137 |   using rsimp_ALTs.simps(3) apply presburger
 | 
|  |    138 |   apply auto
 | 
|  |    139 |   apply(subst rsimp_idem)+
 | 
|  |    140 |   by (metis comp_apply rsimp_idem)
 | 
|  |    141 | 
 | 
|  |    142 | lemma spawn_simp_distinct:
 | 
|  |    143 |   shows "rsimp (rsimp_ALTs (rsa @ (rdistinct rs (set rsa)))) = rsimp (rsimp_ALTs (rsa @ rs))
 | 
|  |    144 | \<and> (a1 \<in> set rsa1 \<longrightarrow> rsimp (rsimp_ALTs (rsa1 @ rs)) = rsimp (rsimp_ALTs (rsa1 @ a1 # rs)))
 | 
|  |    145 | \<and> rsimp  (rsimp_ALTs (rsc @ rs)) = rsimp (rsimp_ALTs (rsc @ (rdistinct rs (set rsc))))"
 | 
|  |    146 |   apply(induct rs arbitrary: rsa rsa1 a1 rsc)
 | 
|  |    147 |    apply simp
 | 
|  |    148 |    apply(subgoal_tac "rsimp (rsimp_ALTs (rsa1 @ [a1])) = rsimp (rsimp_ALTs (rsa1 @ (rdistinct [a1] (set rsa1))))")
 | 
|  |    149 |   prefer 2
 | 
|  |    150 |   
 | 
|  |    151 | 
 | 
|  |    152 | 
 | 
|  |    153 | 
 | 
|  |    154 |   oops
 | 
|  |    155 | 
 | 
|  |    156 | lemma inv_one_derx:
 | 
|  |    157 |   shows " RONE = rder xa r2 \<Longrightarrow> r2 = RCHAR xa"
 | 
|  |    158 |   apply(case_tac r2)
 | 
|  |    159 |        apply simp+
 | 
|  |    160 |   using rrexp.distinct(1) apply presburger
 | 
|  |    161 |     apply (metis rder.simps(5) rrexp.distinct(13) rrexp.simps(20))
 | 
|  |    162 |    apply simp+
 | 
|  |    163 |   done
 | 
|  |    164 | 
 | 
|  |    165 | lemma shape_of_derseq:
 | 
|  |    166 |   shows "rder x (RSEQ r1 r2) = RSEQ (rder x r1) r2 \<or> rder x (RSEQ r1 r2) = (RALT (RSEQ (rder x r1) r2) (rder x r2))"
 | 
|  |    167 |   using rder.simps(5) by presburger
 | 
|  |    168 | lemma shape_of_derseq2:
 | 
|  |    169 |   shows "rder x (RSEQ r11 r12) = RSEQ x41 x42 \<Longrightarrow> x41 = rder x r11"
 | 
|  |    170 |   by (metis rrexp.distinct(25) rrexp.inject(2) shape_of_derseq)
 | 
|  |    171 | 
 | 
|  |    172 | lemma alts_preimage_case1:
 | 
|  |    173 |   shows "rder x r = RALTS [r] \<Longrightarrow> \<exists>ra. r = RALTS [ra]"
 | 
|  |    174 |   apply(case_tac r)
 | 
|  |    175 |        apply simp+
 | 
|  |    176 |   apply (metis rrexp.simps(12) rrexp.simps(20))
 | 
|  |    177 |   apply (metis rrexp.inject(3) rrexp.simps(30) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3) shape_of_derseq)
 | 
|  |    178 |    apply auto[1]
 | 
|  |    179 |   by auto
 | 
|  |    180 | 
 | 
|  |    181 | lemma alts_preimage_case2:
 | 
|  |    182 |   shows "rder x r = RALT r1 r2 \<Longrightarrow> \<exists>ra rb. (r = RSEQ ra rb \<or> r = RALT ra rb)"
 | 
|  |    183 |   apply(case_tac r)
 | 
|  |    184 |        apply simp+
 | 
|  |    185 |   apply (metis rrexp.distinct(15) rrexp.distinct(7))
 | 
|  |    186 |     apply simp
 | 
|  |    187 |   apply auto[1]
 | 
|  |    188 |   by auto
 | 
|  |    189 | 
 | 
|  |    190 | lemma alts_preimage_case2_2:
 | 
|  |    191 |   shows "rder x r = RALT r1 r2 \<Longrightarrow> (\<exists>ra rb. r = RSEQ ra rb) \<or> (\<exists>rc rd. r = RALT rc rd)"
 | 
|  |    192 |   using alts_preimage_case2 by blast
 | 
|  |    193 | 
 | 
|  |    194 | lemma alts_preimage_case3:
 | 
|  |    195 |   shows "rder x r = RALT r1 r2 \<Longrightarrow>  (\<exists>ra rb. r = RSEQ ra rb) \<or> (\<exists>rcs rc rd. r = RALTS rcs \<and> rcs = [rc, rd])"
 | 
|  |    196 |   using alts_preimage_case2 by blast
 | 
|  |    197 | 
 | 
|  |    198 | lemma star_seq:
 | 
|  |    199 |   shows "rder x (RSEQ (RSTAR a) b) = RALT (RSEQ (RSEQ (rder x a) (RSTAR a)) b) (rder x b)"
 | 
|  |    200 |   using rder.simps(5) rder.simps(6) rnullable.simps(6) by presburger
 | 
|  |    201 | 
 | 
|  |    202 | lemma language_equality_id1:
 | 
|  |    203 |   shows "\<not>rnullable a \<Longrightarrow> rder x (RSEQ (RSTAR a) b) = rder x (RALT (RSEQ (RSEQ a (RSTAR a)) b) b)"
 | 
|  |    204 |   apply (subst star_seq)
 | 
|  |    205 |   apply simp
 | 
|  |    206 |   done
 | 
|  |    207 | 
 | 
|  |    208 | 
 | 
|  |    209 | 
 | 
|  |    210 | lemma distinct_der_set:
 | 
|  |    211 |   shows "(rder x) ` rset = dset \<Longrightarrow>
 | 
|  |    212 | rsimp (rsimp_ALTs (map (rder x) (rdistinct rs rset))) = rsimp ( rsimp_ALTs (rdistinct (map (rder x) rs) dset))"
 | 
|  |    213 |   apply(induct rs arbitrary: rset dset)
 | 
|  |    214 |    apply simp
 | 
|  |    215 |   apply(case_tac "a \<in> rset")
 | 
|  |    216 |    apply(subgoal_tac "rder x a \<in> dset")
 | 
|  |    217 |   prefer 2
 | 
|  |    218 |     apply blast
 | 
|  |    219 |    apply simp
 | 
|  |    220 |   apply(case_tac "rder x a \<notin> dset")
 | 
|  |    221 |    prefer 2
 | 
|  |    222 |    apply simp
 | 
|  |    223 |  
 | 
|  |    224 |   oops
 | 
|  |    225 | 
 | 
|  |    226 | lemma map_concat_cons:
 | 
|  |    227 |   shows "map f rsa @ f a # rs = map f (rsa @ [a]) @ rs"
 | 
|  |    228 |   by simp
 | 
|  |    229 | 
 | 
|  |    230 | lemma neg_removal_element_of:
 | 
|  |    231 |   shows " \<not> a \<notin> aset \<Longrightarrow> a \<in> aset"
 | 
|  |    232 |   by simp
 | 
|  |    233 | 
 | 
|  |    234 | lemma simp_more_flts:
 | 
|  |    235 |   shows "rsimp (rsimp_ALTs (rdistinct rs {})) = rsimp (rsimp_ALTs (rdistinct (rflts rs) {}))"
 | 
|  |    236 | 
 | 
|  |    237 |   oops
 | 
|  |    238 | 
 | 
| 465 |    239 | lemma simp_more_distinct1:
 | 
|  |    240 |   shows "rsimp (rsimp_ALTs rs) = rsimp (rsimp_ALTs (rdistinct rs {}))"
 | 
|  |    241 |   apply(induct rs)
 | 
|  |    242 |    apply simp
 | 
|  |    243 |   apply simp
 | 
|  |    244 |   oops
 | 
|  |    245 | 
 | 
|  |    246 | 
 | 
|  |    247 | (*
 | 
|  |    248 | \<and>
 | 
|  |    249 |   rsimp (rsimp_ALTs (rsb @ (rdistinct rs (set rsb)))) = 
 | 
|  |    250 |   rsimp (rsimp_ALTs (rsb @ (rdistinct (rflts rs) (set rsb))))
 | 
|  |    251 | *)
 | 
|  |    252 | lemma simp_removes_duplicate2:
 | 
|  |    253 |   shows "a "
 | 
|  |    254 | 
 | 
|  |    255 |   oops
 | 
|  |    256 | 
 | 
|  |    257 | lemma flts_removes0:
 | 
|  |    258 |   shows "  rflts (rs @ [RZERO])  =
 | 
|  |    259 |            rflts rs"
 | 
|  |    260 |   apply(induct rs)
 | 
|  |    261 |    apply simp
 | 
|  |    262 |   by (metis append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
 | 
|  |    263 |   
 | 
|  |    264 | lemma flts_keeps1:
 | 
|  |    265 |   shows " rflts (rs @ [RONE]) = 
 | 
|  |    266 |           rflts  rs @ [RONE] "
 | 
|  |    267 |   apply (induct rs)
 | 
|  |    268 |    apply simp
 | 
|  |    269 |   by (metis append.assoc append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
 | 
|  |    270 | 
 | 
|  |    271 | lemma flts_keeps_others:
 | 
|  |    272 |   shows "\<lbrakk>a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1\<rbrakk> \<Longrightarrow>rflts (rs @ [a]) = rflts rs @ [a]"
 | 
|  |    273 |   apply(induct rs)
 | 
|  |    274 |    apply simp
 | 
|  |    275 |   apply (simp add: rflts_def_idiot)
 | 
|  |    276 |   apply(case_tac a)
 | 
|  |    277 |        apply simp
 | 
|  |    278 |   using flts_keeps1 apply blast
 | 
|  |    279 |      apply (metis append.assoc append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
 | 
|  |    280 |   apply (metis append.assoc append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
 | 
|  |    281 |   apply blast
 | 
|  |    282 |   by (metis append.assoc append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
 | 
|  |    283 | 
 | 
|  |    284 | 
 | 
|  |    285 | lemma rflts_def_idiot2:
 | 
|  |    286 |   shows "\<lbrakk>a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1; a \<in> set rs\<rbrakk> \<Longrightarrow> a \<in> set (rflts rs)"
 | 
|  |    287 |   apply(induct rs)
 | 
|  |    288 |    apply simp
 | 
|  |    289 |   by (metis append.assoc in_set_conv_decomp insert_iff list.simps(15) rflts.simps(2) rflts.simps(3) rflts_def_idiot)
 | 
|  |    290 | 
 | 
|  |    291 | lemma rflts_spills_last:
 | 
|  |    292 |   shows "a = RALTS rs \<Longrightarrow> rflts (rs1 @ [a]) = rflts rs1 @ rs"
 | 
|  |    293 |   apply (induct rs1)
 | 
|  |    294 |   apply simp
 | 
|  |    295 |   by (metis append.assoc append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
 | 
|  |    296 | 
 | 
|  |    297 | 
 | 
|  |    298 | lemma spilled_alts_contained:
 | 
|  |    299 |   shows "\<lbrakk>a = RALTS rs ; a \<in> set rs1\<rbrakk> \<Longrightarrow> \<forall>r \<in> set rs. r \<in> set (rflts rs1)"
 | 
|  |    300 |   apply(induct rs1)
 | 
|  |    301 |    apply simp 
 | 
|  |    302 |   apply(case_tac "a = aa")
 | 
|  |    303 |    apply simp
 | 
|  |    304 |   apply(subgoal_tac " a \<in> set rs1")
 | 
|  |    305 |   prefer 2
 | 
|  |    306 |    apply (meson set_ConsD)
 | 
|  |    307 |   apply(case_tac aa)
 | 
|  |    308 |   using rflts.simps(2) apply presburger
 | 
|  |    309 |       apply fastforce
 | 
|  |    310 |   apply fastforce
 | 
|  |    311 |   apply fastforce
 | 
|  |    312 |   apply fastforce
 | 
|  |    313 |   by fastforce
 | 
|  |    314 | 
 | 
|  |    315 | lemma distinct_removes_duplicate_flts:
 | 
|  |    316 |   shows " a \<in> set rsa
 | 
|  |    317 |        \<Longrightarrow> rdistinct (rflts (map rsimp rsa @ [rsimp a])) {} =
 | 
|  |    318 |            rdistinct (rflts (map rsimp rsa)) {}"
 | 
|  |    319 |   apply(subgoal_tac "rsimp a \<in> set (map rsimp rsa)")
 | 
|  |    320 |   prefer 2
 | 
|  |    321 |   apply simp
 | 
|  |    322 |   apply(induct "rsimp a")
 | 
|  |    323 |        apply simp
 | 
|  |    324 |   using flts_removes0 apply presburger
 | 
|  |    325 |       apply(subgoal_tac " rdistinct (rflts (map rsimp rsa @ [rsimp a])) {} =  
 | 
|  |    326 |                           rdistinct (rflts (map rsimp rsa @ [RONE])) {}")
 | 
|  |    327 |       apply (simp only:)
 | 
|  |    328 |        apply(subst flts_keeps1)
 | 
|  |    329 |   apply (metis distinct_removes_last2 rflts_def_idiot2 rrexp.simps(20) rrexp.simps(6))
 | 
|  |    330 |       apply presburger
 | 
|  |    331 |         apply(subgoal_tac " rdistinct (rflts (map rsimp rsa @ [rsimp a]))    {} =  
 | 
|  |    332 |                             rdistinct ((rflts (map rsimp rsa)) @ [RCHAR x]) {}")
 | 
|  |    333 |       apply (simp only:)
 | 
|  |    334 |       prefer 2
 | 
|  |    335 |       apply (metis flts_keeps_others rrexp.distinct(21) rrexp.distinct(3))
 | 
|  |    336 |   apply (metis distinct_removes_last2 rflts_def_idiot2 rrexp.distinct(21) rrexp.distinct(3))
 | 
|  |    337 | 
 | 
|  |    338 |     apply (metis distinct_removes_last2 flts_keeps_others rflts_def_idiot2 rrexp.distinct(25) rrexp.distinct(5))
 | 
|  |    339 |    prefer 2
 | 
|  |    340 |    apply (metis distinct_removes_last2 flts_keeps_others flts_removes0 rflts_def_idiot2 rrexp.distinct(29))
 | 
|  |    341 |   apply(subgoal_tac "rflts (map rsimp rsa @ [rsimp a]) = rflts (map rsimp rsa) @ x")
 | 
|  |    342 |   prefer 2
 | 
|  |    343 |   apply (simp add: rflts_spills_last)
 | 
|  |    344 |   apply(simp only:)
 | 
|  |    345 |   apply(subgoal_tac "\<forall> r \<in> set x. r \<in> set (rflts (map rsimp rsa))")
 | 
|  |    346 |   prefer 2
 | 
|  |    347 |   using spilled_alts_contained apply presburger
 | 
|  |    348 |   by (metis append_self_conv distinct_removes_list in_set_conv_decomp rev_exhaust)
 | 
|  |    349 | 
 | 
|  |    350 | lemma flts_middle0:
 | 
|  |    351 |   shows "rflts (rsa @ RZERO # rsb) = rflts (rsa @ rsb)"
 | 
|  |    352 |   apply(induct rsa)
 | 
|  |    353 |   apply simp
 | 
|  |    354 |   by (metis append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
 | 
|  |    355 | 
 | 
|  |    356 | lemma flts_middle01:
 | 
|  |    357 |   shows "rflts (rsa @ [RZERO] @ rsb) = rflts (rsa @ rsb)"
 | 
|  |    358 |   by (simp add: flts_middle0)
 | 
|  |    359 | 
 | 
|  |    360 | lemma flts_append1:
 | 
|  |    361 |   shows "\<lbrakk>a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1; a \<in> set rs\<rbrakk>  \<Longrightarrow>
 | 
|  |    362 |          rflts (rsa @ [a] @ rsb) = rflts rsa @ [a] @ (rflts rsb)"
 | 
|  |    363 |   apply(induct rsa arbitrary: rsb)
 | 
|  |    364 |    apply simp
 | 
|  |    365 |   using rflts_def_idiot apply presburger
 | 
|  |    366 |   apply(case_tac aa)  
 | 
|  |    367 |        apply simp+
 | 
|  |    368 |   done
 | 
|  |    369 | 
 | 
|  |    370 | lemma flts_append:
 | 
|  |    371 |   shows "rflts (rs1 @ rs2) = rflts rs1 @ rflts rs2"
 | 
|  |    372 |   apply(induct rs1)
 | 
|  |    373 |    apply simp
 | 
|  |    374 |   apply(case_tac a)
 | 
|  |    375 |        apply simp+
 | 
|  |    376 |   done
 | 
|  |    377 | 
 | 
|  |    378 | lemma simp_removes_duplicate1:
 | 
|  |    379 |   shows  " a \<in> set rsa \<Longrightarrow> rsimp (RALTS (rsa @ [a])) =  rsimp (RALTS (rsa))"
 | 
|  |    380 | and " rsimp (RALTS (a1 # rsa @ [a1])) = rsimp (RALTS (a1 # rsa))"
 | 
|  |    381 |   apply(induct rsa arbitrary: a1)
 | 
|  |    382 |      apply simp
 | 
|  |    383 |     apply simp
 | 
|  |    384 |   prefer 2
 | 
|  |    385 |   apply(case_tac "a = aa")
 | 
|  |    386 |      apply simp
 | 
|  |    387 |     apply simp
 | 
|  |    388 |   apply (metis Cons_eq_appendI Cons_eq_map_conv distinct_removes_duplicate_flts list.set_intros(2))
 | 
|  |    389 |   apply (metis append_Cons append_Nil distinct_removes_duplicate_flts list.set_intros(1) list.simps(8) list.simps(9))
 | 
|  |    390 |   by (metis (mono_tags, lifting) append_Cons distinct_removes_duplicate_flts list.set_intros(1) list.simps(8) list.simps(9) map_append rsimp.simps(2))
 | 
|  |    391 |   
 | 
|  |    392 | lemma simp_removes_duplicate2:
 | 
|  |    393 |   shows "a \<in> set rsa \<Longrightarrow> rsimp (RALTS (rsa @ [a] @ rsb)) = rsimp (RALTS (rsa @ rsb))"
 | 
|  |    394 |   apply(induct rsb arbitrary: rsa)
 | 
|  |    395 |    apply simp
 | 
|  |    396 |   using distinct_removes_duplicate_flts apply auto[1]
 | 
|  |    397 |   by (metis append.assoc head_one_more_simp rsimp.simps(2) simp_flatten simp_removes_duplicate1(1))
 | 
|  |    398 | 
 | 
|  |    399 | lemma simp_removes_duplicate3:
 | 
|  |    400 |   shows "a \<in> set rsa \<Longrightarrow> rsimp (RALTS (rsa @ a # rsb)) = rsimp (RALTS (rsa @ rsb))"
 | 
|  |    401 |   using simp_removes_duplicate2 by auto
 | 
|  |    402 | 
 | 
|  |    403 | lemma distinct_removes_middle4:
 | 
|  |    404 |   shows "a \<in> set rsa \<Longrightarrow> rdistinct (rsa @ [a] @ rsb) rset = rdistinct (rsa @ rsb) rset"
 | 
|  |    405 |   using distinct_removes_middle(1) by fastforce
 | 
|  |    406 | 
 | 
|  |    407 | lemma distinct_removes_middle_list:
 | 
|  |    408 |   shows "\<forall>a \<in> set x. a \<in> set rsa \<Longrightarrow> rdistinct (rsa @ x @ rsb) rset = rdistinct (rsa @ rsb) rset"
 | 
|  |    409 |   apply(induct x)
 | 
|  |    410 |    apply simp
 | 
|  |    411 |   by (simp add: distinct_removes_middle3)
 | 
|  |    412 | 
 | 
|  |    413 | 
 | 
|  |    414 | lemma distinct_removes_duplicate_flts2:
 | 
|  |    415 |   shows " a \<in> set rsa
 | 
|  |    416 |        \<Longrightarrow> rdistinct (rflts (rsa @ [a] @ rsb)) {} =
 | 
|  |    417 |            rdistinct (rflts (rsa @ rsb)) {}"
 | 
|  |    418 |   apply(induct a arbitrary: rsb)
 | 
|  |    419 |   using flts_middle01 apply presburger
 | 
|  |    420 |       apply(subgoal_tac "rflts (rsa @ [RONE] @ rsb) = rflts rsa @ [RONE] @ rflts rsb")
 | 
|  |    421 |   prefer 2
 | 
|  |    422 |   using flts_append1 apply blast
 | 
|  |    423 |       apply simp
 | 
|  |    424 |       apply(subgoal_tac "RONE \<in> set (rflts rsa)")
 | 
|  |    425 |   prefer 2
 | 
|  |    426 |   using rflts_def_idiot2 apply blast
 | 
|  |    427 |       apply(subst distinct_removes_middle3)
 | 
|  |    428 |        apply simp
 | 
|  |    429 |   using flts_append apply presburger
 | 
|  |    430 |      apply simp
 | 
|  |    431 |   apply (metis distinct_removes_middle3 flts_append in_set_conv_decomp rflts.simps(5))
 | 
|  |    432 |   apply (metis distinct_removes_middle(1) flts_append flts_append1 rflts_def_idiot2 rrexp.distinct(25) rrexp.distinct(5))
 | 
|  |    433 |    apply(subgoal_tac "rflts (rsa @ [RALTS x] @ rsb) = rflts rsa @ x @ rflts rsb")
 | 
|  |    434 |     prefer 2
 | 
|  |    435 |   apply (simp add: flts_append)
 | 
|  |    436 |    apply (simp only:)
 | 
|  |    437 | 
 | 
|  |    438 |    apply(subgoal_tac "\<forall>r1 \<in> set x. r1 \<in> set (rflts rsa)")
 | 
|  |    439 |     prefer 2
 | 
|  |    440 |   using spilled_alts_contained apply blast
 | 
|  |    441 |   apply(subst flts_append)
 | 
|  |    442 |   using distinct_removes_middle_list apply blast
 | 
|  |    443 |   using distinct_removes_middle2 flts_append rflts_def_idiot2 by fastforce
 | 
|  |    444 | 
 | 
|  |    445 | 
 | 
|  |    446 | lemma simp_removes_duplicate:
 | 
|  |    447 |   shows "a \<in> set rsa \<Longrightarrow> rsimp (rsimp_ALTs (rsa @ a # rs)) =  rsimp (rsimp_ALTs (rsa @ rs))"
 | 
|  |    448 |   apply(subgoal_tac "rsimp (rsimp_ALTs (rsa @ a # rs)) = rsimp (RALTS (rsa @ a # rs))")
 | 
|  |    449 |    prefer 2
 | 
|  |    450 |   apply (smt (verit, best) Cons_eq_append_conv append_is_Nil_conv empty_set equals0D list.distinct(1) rsimp_ALTs.elims)
 | 
|  |    451 |   apply(simp only:)
 | 
|  |    452 |   apply simp
 | 
|  |    453 |   apply(subgoal_tac "(rdistinct (rflts (map rsimp rsa @ rsimp a # map rsimp rs)) {}) = (rdistinct (rflts (map rsimp rsa @  map rsimp rs)) {})")
 | 
|  |    454 |    apply(simp only:)
 | 
|  |    455 |   prefer 2
 | 
|  |    456 |    apply(subgoal_tac "rsimp a \<in> set (map rsimp rsa)")
 | 
|  |    457 |     prefer 2
 | 
|  |    458 |   apply simp
 | 
|  |    459 |   using distinct_removes_duplicate_flts2 apply force
 | 
|  |    460 |   apply(case_tac rsa)
 | 
|  |    461 |    apply simp
 | 
|  |    462 |   apply(case_tac rs)
 | 
|  |    463 |    apply simp
 | 
|  |    464 |    apply(case_tac list)
 | 
|  |    465 |     apply simp
 | 
|  |    466 |   using idem_after_simp1 apply presburger
 | 
|  |    467 |    apply simp+
 | 
|  |    468 |   apply(subgoal_tac "rsimp_ALTs (aa # list @ aaa # lista) = RALTS (aa # list @ aaa # lista)")
 | 
|  |    469 |    apply simp
 | 
|  |    470 |   using rsimpalts_conscons by presburger
 | 
| 467 |    471 | 
 | 
| 468 |    472 | lemma no0_flts:
 | 
|  |    473 |   shows "RZERO \<notin> set (rflts rs)"
 | 
|  |    474 |   apply (induct rs)
 | 
|  |    475 |    apply simp
 | 
|  |    476 |   apply(case_tac a)
 | 
|  |    477 |        apply simp+
 | 
|  |    478 |   oops
 | 
| 467 |    479 | 
 | 
|  |    480 | 
 | 
|  |    481 | 
 | 
|  |    482 | lemma distinct_flts_no0:
 | 
| 468 |    483 |   shows "  rflts (map rsimp (rdistinct rs (insert RZERO rset)))  =
 | 
|  |    484 |            rflts (map rsimp (rdistinct rs rset))  "
 | 
|  |    485 |   
 | 
|  |    486 |   apply(induct rs arbitrary: rset)
 | 
| 467 |    487 |    apply simp
 | 
|  |    488 |   apply(case_tac a)
 | 
| 468 |    489 |   apply simp+
 | 
|  |    490 |     apply (smt (verit, ccfv_SIG) rflts.simps(2) rflts.simps(3) rflts_def_idiot)
 | 
|  |    491 |   prefer 2
 | 
|  |    492 |   apply simp  
 | 
|  |    493 |   by (smt (verit, ccfv_threshold) Un_insert_right insert_iff list.simps(9) rdistinct.simps(2) rflts.simps(2) rflts.simps(3) rflts_def_idiot rrexp.distinct(7))
 | 
|  |    494 | 
 | 
| 467 |    495 | 
 | 
|  |    496 | 
 | 
|  |    497 | lemma simp_der_flts:
 | 
|  |    498 |   shows "rsimp (RALTS (rdistinct (map (rder x) (rflts rs)) rset))= 
 | 
|  |    499 |          rsimp (RALTS (rdistinct (rflts (map (rder x) rs)) rset))"
 | 
|  |    500 | 
 | 
| 468 |    501 |   apply(induct rs)
 | 
| 467 |    502 |    apply simp
 | 
|  |    503 |   apply(case_tac a)
 | 
|  |    504 |        apply simp
 | 
|  |    505 |       apply(case_tac "RZERO \<in> rset")
 | 
| 468 |    506 |        apply simp+
 | 
|  |    507 |   using distinct_flts_no0 apply presburger
 | 
|  |    508 |      apply (case_tac "x = x3")
 | 
|  |    509 |   prefer 2
 | 
|  |    510 |   apply simp
 | 
|  |    511 |   using distinct_flts_no0 apply presburger
 | 
|  |    512 |   apply(case_tac "RONE \<in> rset")
 | 
|  |    513 |      apply simp+
 | 
| 465 |    514 |   
 | 
| 467 |    515 |   sorry
 | 
|  |    516 | 
 | 
|  |    517 | 
 | 
| 465 |    518 | lemma simp_der_pierce_flts:
 | 
|  |    519 |   shows " rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {})) =
 | 
|  |    520 |            rsimp (rsimp_ALTs (rdistinct (rflts (map (rder x) (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {}))"
 | 
| 467 |    521 |   sorry
 | 
|  |    522 | 
 | 
| 465 |    523 | 
 | 
| 453 |    524 | 
 | 
|  |    525 | 
 | 
|  |    526 | lemma simp_more_distinct:
 | 
| 465 |    527 |   shows "rsimp  (rsimp_ALTs (rsa @ rs)) = rsimp (rsimp_ALTs (rsa @ (rdistinct rs (set rsa)))) "
 | 
| 467 |    528 |   
 | 
|  |    529 | 
 | 
| 465 |    530 | 
 | 
| 453 |    531 | 
 | 
|  |    532 |   sorry
 | 
|  |    533 | 
 | 
|  |    534 | lemma non_empty_list:
 | 
|  |    535 |   shows "a \<in> set as \<Longrightarrow> as \<noteq> []"
 | 
|  |    536 |   by (metis empty_iff empty_set)
 | 
|  |    537 | 
 | 
| 456 |    538 | lemma distinct_comp:
 | 
|  |    539 |   shows "rdistinct (rs1@rs2) {} = (rdistinct rs1 {}) @ (rdistinct rs2 (set rs1))"
 | 
|  |    540 |   apply(induct rs2 arbitrary: rs1)
 | 
|  |    541 |    apply simp
 | 
|  |    542 |   apply(subgoal_tac "rs1 @ a # rs2 = (rs1 @ [a]) @ rs2")
 | 
|  |    543 |    apply(simp only:)
 | 
|  |    544 |    apply(case_tac "a \<in> set rs1")
 | 
|  |    545 |     apply simp
 | 
|  |    546 |   oops
 | 
| 453 |    547 | 
 | 
| 456 |    548 | lemma instantiate1:
 | 
|  |    549 |   shows "\<lbrakk>\<And>ab rset1.  rdistinct (ab # as) rset1 = rdistinct (ab # as @ [ab]) rset1\<rbrakk> \<Longrightarrow>  
 | 
|  |    550 | rdistinct (aa # as) rset = rdistinct (aa # as @ [aa]) rset"
 | 
|  |    551 |   apply(drule_tac x = "aa" in meta_spec)
 | 
|  |    552 |   apply(drule_tac x = "rset" in meta_spec)
 | 
| 453 |    553 |   apply simp
 | 
| 456 |    554 |   done
 | 
|  |    555 |   
 | 
|  |    556 | 
 | 
|  |    557 | lemma not_head_elem:
 | 
|  |    558 |   shows " \<lbrakk>aa \<in> set (a # as); aa \<notin> (set as)\<rbrakk> \<Longrightarrow> a = aa"
 | 
|  |    559 |   
 | 
|  |    560 |   by fastforce
 | 
|  |    561 | 
 | 
|  |    562 | (*
 | 
|  |    563 |   apply simp
 | 
|  |    564 |   apply (metis append_Cons)
 | 
|  |    565 |   apply(case_tac "ab \<in> rset1")
 | 
|  |    566 |   apply (metis (no_types, opaque_lifting) Un_insert_left append_Cons insert_iff rdistinct.simps(2) sup_bot_left)
 | 
|  |    567 |   apply(subgoal_tac "rdistinct (ab # (aa # as) @ [ab]) rset1 = 
 | 
|  |    568 |                ab # (rdistinct ((aa # as) @ [ab]) (insert ab rset1))")
 | 
|  |    569 |    apply(simp only:)
 | 
|  |    570 |    apply(subgoal_tac "rdistinct (ab # aa # as) rset1 = ab # (rdistinct (aa # as) (insert ab rset1))")
 | 
|  |    571 |   apply(simp only:)
 | 
|  |    572 |     apply(subgoal_tac "rdistinct ((aa # as) @ [ab]) (insert ab rset1) = rdistinct (aa # as) (insert ab rset1)")
 | 
|  |    573 |   apply blast
 | 
|  |    574 | *)
 | 
|  |    575 |   
 | 
| 453 |    576 | 
 | 
|  |    577 | lemma flts_identity1:
 | 
|  |    578 |   shows  "rflts (rs @ [RONE]) = rflts rs @ [RONE] "
 | 
|  |    579 |   apply(induct rs)
 | 
|  |    580 |    apply simp+
 | 
|  |    581 |   apply(case_tac a)
 | 
|  |    582 |        apply simp
 | 
|  |    583 |       apply simp+
 | 
|  |    584 |   done
 | 
|  |    585 | 
 | 
|  |    586 | lemma flts_identity10:
 | 
|  |    587 |   shows " rflts (rs @ [RCHAR c]) = rflts rs @ [RCHAR c]"
 | 
|  |    588 |   apply(induct rs)
 | 
|  |    589 |    apply simp+
 | 
|  |    590 |   apply(case_tac a)
 | 
|  |    591 |        apply simp+
 | 
|  |    592 |   done
 | 
|  |    593 | 
 | 
|  |    594 | lemma flts_identity11:
 | 
|  |    595 |   shows " rflts (rs @ [RSEQ r1 r2]) = rflts rs @ [RSEQ r1 r2]"
 | 
|  |    596 |   apply(induct rs)
 | 
|  |    597 |    apply simp+
 | 
|  |    598 |   apply(case_tac a)
 | 
|  |    599 |        apply simp+
 | 
|  |    600 |   done
 | 
|  |    601 | 
 | 
|  |    602 | lemma flts_identity12:
 | 
|  |    603 |   shows " rflts (rs @ [RSTAR r0]) = rflts rs @ [RSTAR r0]"
 | 
|  |    604 |   apply(induct rs)
 | 
|  |    605 |    apply simp+
 | 
|  |    606 |   apply(case_tac a)
 | 
|  |    607 |        apply simp+
 | 
|  |    608 |   done
 | 
|  |    609 | 
 | 
|  |    610 | lemma flts_identity2:
 | 
|  |    611 |   shows "a \<noteq> RZERO \<and> (\<forall>rs. a \<noteq> RALTS rs) \<Longrightarrow>  rflts (rs @ [a]) = rflts rs @ [a]"
 | 
|  |    612 |   apply(case_tac a)
 | 
|  |    613 |        apply simp
 | 
|  |    614 |   using flts_identity1 apply auto[1]
 | 
|  |    615 |   using flts_identity10 apply blast
 | 
|  |    616 |   using flts_identity11 apply auto[1]
 | 
|  |    617 |    apply blast
 | 
|  |    618 |   using flts_identity12 by presburger
 | 
| 456 |    619 | 
 | 
|  |    620 | lemma flts_identity3:
 | 
|  |    621 |   shows "a = RZERO  \<Longrightarrow> rflts (rs @ [a]) = rflts rs"
 | 
|  |    622 |   apply simp
 | 
|  |    623 |   apply(induct rs)
 | 
|  |    624 |    apply simp+
 | 
|  |    625 |   apply(case_tac aa)
 | 
|  |    626 |        apply simp+
 | 
|  |    627 |   done
 | 
|  |    628 | 
 | 
|  |    629 | lemma distinct_removes_last3:
 | 
| 465 |    630 |   shows "\<lbrakk>a \<in> set as\<rbrakk>
 | 
| 456 |    631 |     \<Longrightarrow> rdistinct as {} = rdistinct (as @ [a]) {}"
 | 
| 465 |    632 |   by (simp add: distinct_removes_last2)
 | 
| 456 |    633 | 
 | 
|  |    634 | lemma set_inclusion_with_flts1:
 | 
|  |    635 |   shows " \<lbrakk>RONE \<in> set rs\<rbrakk> \<Longrightarrow> RONE  \<in> set (rflts rs)"
 | 
|  |    636 |   apply(induct rs)
 | 
|  |    637 |    apply simp
 | 
|  |    638 |   apply(case_tac " RONE \<in> set rs")
 | 
|  |    639 |    apply simp
 | 
|  |    640 |   apply (metis Un_upper2 insert_absorb insert_subset list.set_intros(2) rflts.simps(2) rflts.simps(3) rflts_def_idiot set_append)
 | 
|  |    641 |   apply(case_tac "RONE = a")
 | 
|  |    642 |    apply simp
 | 
|  |    643 |   apply simp
 | 
|  |    644 |   done
 | 
|  |    645 | 
 | 
|  |    646 | lemma set_inclusion_with_flts10:
 | 
|  |    647 |   shows " \<lbrakk>RCHAR x \<in> set rs\<rbrakk> \<Longrightarrow> RCHAR x  \<in> set (rflts rs)"
 | 
|  |    648 |   apply(induct rs)
 | 
|  |    649 |    apply simp
 | 
|  |    650 |   apply(case_tac " RCHAR x \<in> set rs")
 | 
|  |    651 |    apply simp
 | 
|  |    652 |   apply (metis Un_upper2 insert_absorb insert_subset rflts.simps(2) rflts.simps(3) rflts_def_idiot set_append set_subset_Cons)
 | 
|  |    653 |   apply(case_tac "RCHAR x = a")
 | 
|  |    654 |    apply simp
 | 
|  |    655 |   apply fastforce
 | 
|  |    656 |   apply simp
 | 
|  |    657 |   done
 | 
|  |    658 | 
 | 
|  |    659 | lemma set_inclusion_with_flts11:
 | 
|  |    660 |   shows " \<lbrakk>RSEQ r1 r2 \<in> set rs\<rbrakk> \<Longrightarrow> RSEQ r1 r2  \<in> set (rflts rs)"
 | 
|  |    661 |   apply(induct rs)
 | 
|  |    662 |    apply simp
 | 
|  |    663 |   apply(case_tac " RSEQ r1 r2 \<in> set rs")
 | 
|  |    664 |    apply simp
 | 
|  |    665 |   apply (metis Un_upper2 insert_absorb insert_subset rflts.simps(2) rflts.simps(3) rflts_def_idiot set_append set_subset_Cons)
 | 
|  |    666 |   apply(case_tac "RSEQ r1 r2 = a")
 | 
|  |    667 |    apply simp
 | 
|  |    668 |   apply fastforce
 | 
|  |    669 |   apply simp
 | 
|  |    670 |   done
 | 
|  |    671 | 
 | 
|  |    672 | 
 | 
|  |    673 | lemma set_inclusion_with_flts:
 | 
|  |    674 |   shows " \<lbrakk>a \<in> set as; rsimp a \<in> set (map rsimp as); rsimp a = RONE\<rbrakk> \<Longrightarrow> rsimp a \<in> set (rflts (map rsimp as))"
 | 
|  |    675 |   by (simp add: set_inclusion_with_flts1)
 | 
| 453 |    676 |   
 | 
| 456 |    677 | lemma "\<And>x5. \<lbrakk>a \<in> set as; rsimp a \<in> set (map rsimp as); rsimp a = RALTS x5\<rbrakk>
 | 
|  |    678 |           \<Longrightarrow> rsimp_ALTs (rdistinct (rflts (map rsimp as @ [rsimp a])) {}) = 
 | 
|  |    679 | rsimp_ALTs (rdistinct (rflts (map rsimp as @ x5)) {})"
 | 
|  |    680 | 
 | 
| 465 |    681 |   sorry
 | 
|  |    682 | 
 | 
| 453 |    683 | 
 | 
|  |    684 | lemma last_elem_dup1:
 | 
|  |    685 |   shows " a \<in> set as \<Longrightarrow> rsimp (RALTS (as @ [a] )) = rsimp (RALTS (as ))"
 | 
|  |    686 |   apply simp
 | 
|  |    687 |   apply(subgoal_tac "rsimp a \<in> set (map rsimp as)")
 | 
|  |    688 |   prefer 2
 | 
|  |    689 |    apply simp
 | 
| 456 |    690 |   apply(case_tac "rsimp a")
 | 
|  |    691 |        apply simp
 | 
|  |    692 |   
 | 
|  |    693 |   using flts_identity3 apply presburger
 | 
|  |    694 |       apply(subst flts_identity2)
 | 
|  |    695 |   using rrexp.distinct(1) rrexp.distinct(15) apply presburger
 | 
|  |    696 |       apply(subst distinct_removes_last3[symmetric])
 | 
|  |    697 |   using set_inclusion_with_flts apply blast
 | 
|  |    698 |   apply simp
 | 
|  |    699 |   apply (metis distinct_removes_last3 flts_identity10 set_inclusion_with_flts10)
 | 
|  |    700 |   apply (metis distinct_removes_last3 flts_identity11 set_inclusion_with_flts11)
 | 
| 453 |    701 |   sorry
 | 
|  |    702 | 
 | 
|  |    703 | lemma last_elem_dup:
 | 
|  |    704 |   shows " a \<in> set as \<Longrightarrow> rsimp (rsimp_ALTs (as @ [a] )) = rsimp (rsimp_ALTs (as ))"
 | 
|  |    705 |   apply(induct as rule: rev_induct)
 | 
|  |    706 |    apply simp
 | 
|  |    707 |   apply simp
 | 
|  |    708 |   apply(subgoal_tac "xs \<noteq> []")
 | 
|  |    709 |   prefer 2
 | 
|  |    710 |   
 | 
|  |    711 | 
 | 
|  |    712 |   
 | 
|  |    713 | 
 | 
|  |    714 |   sorry
 | 
|  |    715 | 
 | 
|  |    716 | lemma appeared_before_remove_later:
 | 
|  |    717 |   shows "a \<in>  set as \<Longrightarrow> rsimp (rsimp_ALTs ( as @ a # rs)) = rsimp (rsimp_ALTs (as @ rs))"
 | 
|  |    718 | and "a \<in> set as \<Longrightarrow> rsimp (rsimp_ALTs as ) = rsimp (rsimp_ALTs (as @ [a]))"
 | 
|  |    719 |   apply(induct rs arbitrary: as)
 | 
|  |    720 |    apply simp
 | 
|  |    721 |   
 | 
|  |    722 | 
 | 
|  |    723 |   sorry
 | 
|  |    724 | 
 | 
|  |    725 | lemma distinct_remove_later:
 | 
|  |    726 |   shows "\<lbrakk>rder x a \<in> rder x ` set rsa\<rbrakk>
 | 
|  |    727 |        \<Longrightarrow> rsimp (rsimp_ALTs (map (rder x) rsa @ rder x a # map (rder x) (rdistinct rs (insert a (set rsa))))) =
 | 
|  |    728 |            rsimp (rsimp_ALTs (map (rder x) rsa @ map (rder x) (rdistinct rs (set rsa))))"
 | 
| 451 |    729 |   
 | 
|  |    730 |   sorry
 | 
|  |    731 | 
 | 
|  |    732 | 
 | 
| 453 |    733 | lemma distinct_der_general:
 | 
|  |    734 |   shows "rsimp (rsimp_ALTs (map (rder x) (rsa @ (rdistinct rs (set rsa))))) =
 | 
|  |    735 |  rsimp ( rsimp_ALTs ((map (rder x) rsa)@(rdistinct (map (rder x) rs) (set (map (rder x) rsa)))) )"
 | 
|  |    736 |   apply(induct rs arbitrary: rsa)
 | 
|  |    737 |    apply simp
 | 
|  |    738 |   apply(case_tac "a \<in> set rsa")
 | 
|  |    739 |    apply(subgoal_tac "rder x a \<in> set (map (rder x) rsa)")
 | 
|  |    740 |   apply simp
 | 
|  |    741 |    apply simp
 | 
|  |    742 |   apply(case_tac "rder x a \<notin> set (map (rder x) rsa)")
 | 
|  |    743 |    apply(simp)
 | 
|  |    744 |   apply(subst map_concat_cons)+
 | 
|  |    745 |   apply(drule_tac x = "rsa @ [a]" in meta_spec)
 | 
|  |    746 |    apply simp
 | 
|  |    747 |   apply(drule neg_removal_element_of)
 | 
|  |    748 |   apply simp
 | 
|  |    749 |   apply(subst distinct_remove_later)
 | 
|  |    750 |    apply simp
 | 
|  |    751 |   apply(drule_tac x = "rsa" in meta_spec)
 | 
|  |    752 |   by blast
 | 
|  |    753 | 
 | 
|  |    754 |   
 | 
|  |    755 | 
 | 
|  |    756 | 
 | 
| 451 |    757 | lemma distinct_der:
 | 
|  |    758 |   shows "rsimp (rsimp_ALTs (map (rder x) (rdistinct rs {}))) = rsimp ( rsimp_ALTs (rdistinct (map (rder x) rs) {}))"
 | 
| 453 |    759 |   by (metis distinct_der_general list.simps(8) self_append_conv2 set_empty)
 | 
| 451 |    760 | 
 | 
| 453 |    761 |   
 | 
|  |    762 | 
 | 
|  |    763 | 
 | 
|  |    764 | lemma rders_simp_lambda:
 | 
|  |    765 |   shows " rsimp \<circ> rder x \<circ> (\<lambda>r. rders_simp r xs) = (\<lambda>r. rders_simp r (xs @ [x]))"
 | 
|  |    766 |   using rders_simp_append by auto
 | 
| 451 |    767 | 
 | 
| 453 |    768 | lemma rders_simp_nonempty_simped:
 | 
|  |    769 |   shows "xs \<noteq> [] \<Longrightarrow> rsimp \<circ> (\<lambda>r. rders_simp r xs) = (\<lambda>r. rders_simp r xs)"
 | 
|  |    770 |   using rders_simp_same_simpders rsimp_idem by auto
 | 
|  |    771 | 
 | 
|  |    772 | lemma repeated_altssimp:
 | 
|  |    773 |   shows "\<forall>r \<in> set rs. rsimp r = r \<Longrightarrow> rsimp (rsimp_ALTs (rdistinct (rflts rs) {})) =
 | 
|  |    774 |            rsimp_ALTs (rdistinct (rflts rs) {})"
 | 
|  |    775 |   by (metis map_idI rsimp.simps(2) rsimp_idem)
 | 
| 451 |    776 | 
 | 
| 465 |    777 | 
 | 
|  |    778 | lemma add0_isomorphic:
 | 
|  |    779 |   shows "rsimp_ALTs (rdistinct (rflts [rsimp r, RZERO]) {}) = rsimp r"
 | 
|  |    780 |   sorry
 | 
|  |    781 | 
 | 
|  |    782 | 
 | 
|  |    783 | lemma distinct_append_simp:
 | 
|  |    784 |   shows " rsimp (rsimp_ALTs rs1) = rsimp (rsimp_ALTs rs2) \<Longrightarrow>
 | 
|  |    785 |            rsimp (rsimp_ALTs (f a # rs1)) =
 | 
|  |    786 |            rsimp (rsimp_ALTs (f a # rs2))"
 | 
|  |    787 |   apply(case_tac rs1)
 | 
|  |    788 |    apply simp
 | 
|  |    789 |    apply(case_tac rs2)
 | 
|  |    790 |     apply simp
 | 
|  |    791 |    apply simp
 | 
|  |    792 |    prefer 2
 | 
|  |    793 |    apply(case_tac list)
 | 
|  |    794 |     apply(case_tac rs2)
 | 
|  |    795 |      apply simp
 | 
|  |    796 |   using add0_isomorphic apply blast 
 | 
|  |    797 |     apply simp
 | 
| 467 |    798 |   oops
 | 
| 465 |    799 | 
 | 
| 444 |    800 | lemma alts_closed_form: shows 
 | 
|  |    801 | "rsimp (rders_simp (RALTS rs) s) = 
 | 
|  |    802 | rsimp (RALTS (map (\<lambda>r. rders_simp r s) rs))"
 | 
|  |    803 |   apply(induct s rule: rev_induct)
 | 
|  |    804 |    apply simp
 | 
|  |    805 |   apply simp
 | 
|  |    806 |   apply(subst rders_simp_append)
 | 
|  |    807 |   apply(subgoal_tac " rsimp (rders_simp (rders_simp (RALTS rs) xs) [x]) = 
 | 
|  |    808 |  rsimp(rders_simp (rsimp_ALTs (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {})) [x])")
 | 
|  |    809 |    prefer 2
 | 
|  |    810 |   apply (metis inside_simp_removal rders_simp_one_char)
 | 
|  |    811 |   apply(simp only: )
 | 
| 451 |    812 |   apply(subst rders_simp_one_char)
 | 
|  |    813 |   apply(subst rsimp_idem)
 | 
|  |    814 |   apply(subgoal_tac "rsimp (rder x (rsimp_ALTs (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {}))) =
 | 
|  |    815 |                      rsimp ((rsimp_ALTs (map (rder x) (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {})))) ")
 | 
|  |    816 |   prefer 2
 | 
|  |    817 |   using rder_rsimp_ALTs_commute apply presburger
 | 
|  |    818 |   apply(simp only:)
 | 
|  |    819 |   apply(subgoal_tac "rsimp (rsimp_ALTs (map (rder x) (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {})))
 | 
|  |    820 | = rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {}))")
 | 
|  |    821 |    prefer 2
 | 
|  |    822 |   
 | 
|  |    823 |   using distinct_der apply presburger
 | 
|  |    824 |   apply(simp only:)
 | 
| 453 |    825 |   apply(subgoal_tac " rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {})) =
 | 
|  |    826 |                       rsimp (rsimp_ALTs (rdistinct ( (rflts (map (rder x) (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)))) {}))")
 | 
|  |    827 |    apply(simp only:)
 | 
|  |    828 |   apply(subgoal_tac " rsimp (rsimp_ALTs (rdistinct (rflts (map (rder x) (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {})) = 
 | 
|  |    829 |                       rsimp (rsimp_ALTs (rdistinct (rflts ( (map (rsimp \<circ> (rder x) \<circ> (\<lambda>r. rders_simp r xs)) rs))) {}))")
 | 
|  |    830 |     apply(simp only:)
 | 
|  |    831 |   apply(subst rders_simp_lambda)
 | 
|  |    832 |     apply(subst rders_simp_nonempty_simped)
 | 
|  |    833 |      apply simp
 | 
|  |    834 |     apply(subgoal_tac "\<forall>r \<in> set  (map (\<lambda>r. rders_simp r (xs @ [x])) rs). rsimp r = r")
 | 
|  |    835 |   prefer 2
 | 
|  |    836 |      apply (simp add: rders_simp_same_simpders rsimp_idem)
 | 
|  |    837 |     apply(subst repeated_altssimp)
 | 
|  |    838 |      apply simp
 | 
|  |    839 |   apply fastforce
 | 
| 465 |    840 |    apply (metis inside_simp_removal list.map_comp rder.simps(4) rsimp.simps(2) rsimp_idem)
 | 
|  |    841 |   sledgehammer
 | 
|  |    842 |  (* by (metis inside_simp_removal rder_rsimp_ALTs_commute self_append_conv2 set_empty simp_more_distinct)
 | 
| 451 |    843 | 
 | 
| 465 |    844 |   *)
 | 
| 443 |    845 | 
 | 
| 444 |    846 | lemma alts_closed_form_variant: shows 
 | 
|  |    847 | "s \<noteq> [] \<Longrightarrow> rders_simp (RALTS rs) s = 
 | 
|  |    848 | rsimp (RALTS (map (\<lambda>r. rders_simp r s) rs))"
 | 
|  |    849 |   sorry
 | 
| 443 |    850 | 
 | 
|  |    851 | 
 | 
|  |    852 | 
 | 
| 444 |    853 | lemma star_closed_form:
 | 
|  |    854 |   shows "rders_simp (RSTAR r0) (c#s) = 
 | 
|  |    855 | rsimp ( RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) ))"
 | 
|  |    856 |   apply(induct s)
 | 
|  |    857 |    apply simp
 | 
|  |    858 |   sorry
 | 
| 443 |    859 | 
 | 
|  |    860 | 
 | 
|  |    861 | 
 | 
|  |    862 | lemma seq_closed_form: shows 
 | 
|  |    863 | "rsimp (rders_simp (RSEQ r1 r2) s) = 
 | 
|  |    864 | rsimp ( RALTS ( (RSEQ (rders_simp r1 s) r2) # 
 | 
| 445 |    865 |                 (map (rders_simp r2) (vsuf s r1)) 
 | 
| 443 |    866 |               )  
 | 
|  |    867 |       )"
 | 
|  |    868 |   apply(induct s)
 | 
|  |    869 |   apply simp
 | 
|  |    870 |   sorry
 | 
|  |    871 | 
 | 
|  |    872 | 
 | 
| 444 |    873 | lemma seq_closed_form_variant: shows
 | 
|  |    874 | "s \<noteq> [] \<Longrightarrow> (rders_simp (RSEQ r1 r2) s) = 
 | 
|  |    875 | rsimp (RALTS ((RSEQ (rders_simp r1 s) r2) # (map (rders_simp r2) (vsuf s r1))))"
 | 
| 445 |    876 |   apply(induct s rule: rev_induct)
 | 
|  |    877 |    apply simp
 | 
|  |    878 |   apply(subst rders_simp_append)
 | 
|  |    879 |   apply(subst rders_simp_one_char)
 | 
|  |    880 |   apply(subst rsimp_idem[symmetric])
 | 
|  |    881 |   apply(subst rders_simp_one_char[symmetric])
 | 
|  |    882 |   apply(subst rders_simp_append[symmetric])
 | 
|  |    883 |   apply(insert seq_closed_form)
 | 
|  |    884 |   apply(subgoal_tac "rsimp (rders_simp (RSEQ r1 r2) (xs @ [x]))
 | 
|  |    885 |  = rsimp (RALTS (RSEQ (rders_simp r1 (xs @ [x])) r2 # map (rders_simp r2) (vsuf (xs @ [x]) r1)))")
 | 
|  |    886 |    apply force
 | 
|  |    887 |   by presburger
 | 
| 443 |    888 | 
 | 
| 444 |    889 | end |