thys2/ClosedForms.thy
author Chengsong
Wed, 13 Apr 2022 22:20:08 +0100
changeset 488 370dae790b30
parent 487 9f3d6f09b093
child 489 2b5b3f83e2b6
permissions -rw-r--r--
more sorrys fileld
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
     1
theory ClosedForms imports
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
     2
"BasicIdentities"
443
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
     3
begin
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
     4
453
Chengsong
parents: 451
diff changeset
     5
lemma map_concat_cons:
Chengsong
parents: 451
diff changeset
     6
  shows "map f rsa @ f a # rs = map f (rsa @ [a]) @ rs"
Chengsong
parents: 451
diff changeset
     7
  by simp
Chengsong
parents: 451
diff changeset
     8
Chengsong
parents: 451
diff changeset
     9
lemma neg_removal_element_of:
Chengsong
parents: 451
diff changeset
    10
  shows " \<not> a \<notin> aset \<Longrightarrow> a \<in> aset"
Chengsong
parents: 451
diff changeset
    11
  by simp
Chengsong
parents: 451
diff changeset
    12
465
Chengsong
parents: 456
diff changeset
    13
Chengsong
parents: 456
diff changeset
    14
Chengsong
parents: 456
diff changeset
    15
Chengsong
parents: 456
diff changeset
    16
478
Chengsong
parents: 476
diff changeset
    17
465
Chengsong
parents: 456
diff changeset
    18
Chengsong
parents: 456
diff changeset
    19
478
Chengsong
parents: 476
diff changeset
    20
465
Chengsong
parents: 456
diff changeset
    21
Chengsong
parents: 456
diff changeset
    22
lemma flts_middle0:
Chengsong
parents: 456
diff changeset
    23
  shows "rflts (rsa @ RZERO # rsb) = rflts (rsa @ rsb)"
Chengsong
parents: 456
diff changeset
    24
  apply(induct rsa)
Chengsong
parents: 456
diff changeset
    25
  apply simp
Chengsong
parents: 456
diff changeset
    26
  by (metis append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
Chengsong
parents: 456
diff changeset
    27
Chengsong
parents: 456
diff changeset
    28
lemma flts_middle01:
Chengsong
parents: 456
diff changeset
    29
  shows "rflts (rsa @ [RZERO] @ rsb) = rflts (rsa @ rsb)"
Chengsong
parents: 456
diff changeset
    30
  by (simp add: flts_middle0)
Chengsong
parents: 456
diff changeset
    31
Chengsong
parents: 456
diff changeset
    32
lemma flts_append1:
Chengsong
parents: 456
diff changeset
    33
  shows "\<lbrakk>a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1; a \<in> set rs\<rbrakk>  \<Longrightarrow>
Chengsong
parents: 456
diff changeset
    34
         rflts (rsa @ [a] @ rsb) = rflts rsa @ [a] @ (rflts rsb)"
Chengsong
parents: 456
diff changeset
    35
  apply(induct rsa arbitrary: rsb)
Chengsong
parents: 456
diff changeset
    36
   apply simp
Chengsong
parents: 456
diff changeset
    37
  using rflts_def_idiot apply presburger
Chengsong
parents: 456
diff changeset
    38
  apply(case_tac aa)  
Chengsong
parents: 456
diff changeset
    39
       apply simp+
Chengsong
parents: 456
diff changeset
    40
  done
Chengsong
parents: 456
diff changeset
    41
478
Chengsong
parents: 476
diff changeset
    42
465
Chengsong
parents: 456
diff changeset
    43
Chengsong
parents: 456
diff changeset
    44
lemma simp_removes_duplicate1:
Chengsong
parents: 456
diff changeset
    45
  shows  " a \<in> set rsa \<Longrightarrow> rsimp (RALTS (rsa @ [a])) =  rsimp (RALTS (rsa))"
Chengsong
parents: 456
diff changeset
    46
and " rsimp (RALTS (a1 # rsa @ [a1])) = rsimp (RALTS (a1 # rsa))"
Chengsong
parents: 456
diff changeset
    47
  apply(induct rsa arbitrary: a1)
Chengsong
parents: 456
diff changeset
    48
     apply simp
Chengsong
parents: 456
diff changeset
    49
    apply simp
Chengsong
parents: 456
diff changeset
    50
  prefer 2
Chengsong
parents: 456
diff changeset
    51
  apply(case_tac "a = aa")
Chengsong
parents: 456
diff changeset
    52
     apply simp
Chengsong
parents: 456
diff changeset
    53
    apply simp
Chengsong
parents: 456
diff changeset
    54
  apply (metis Cons_eq_appendI Cons_eq_map_conv distinct_removes_duplicate_flts list.set_intros(2))
Chengsong
parents: 456
diff changeset
    55
  apply (metis append_Cons append_Nil distinct_removes_duplicate_flts list.set_intros(1) list.simps(8) list.simps(9))
Chengsong
parents: 456
diff changeset
    56
  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))
Chengsong
parents: 456
diff changeset
    57
  
Chengsong
parents: 456
diff changeset
    58
lemma simp_removes_duplicate2:
Chengsong
parents: 456
diff changeset
    59
  shows "a \<in> set rsa \<Longrightarrow> rsimp (RALTS (rsa @ [a] @ rsb)) = rsimp (RALTS (rsa @ rsb))"
Chengsong
parents: 456
diff changeset
    60
  apply(induct rsb arbitrary: rsa)
Chengsong
parents: 456
diff changeset
    61
   apply simp
Chengsong
parents: 456
diff changeset
    62
  using distinct_removes_duplicate_flts apply auto[1]
Chengsong
parents: 456
diff changeset
    63
  by (metis append.assoc head_one_more_simp rsimp.simps(2) simp_flatten simp_removes_duplicate1(1))
Chengsong
parents: 456
diff changeset
    64
Chengsong
parents: 456
diff changeset
    65
lemma simp_removes_duplicate3:
Chengsong
parents: 456
diff changeset
    66
  shows "a \<in> set rsa \<Longrightarrow> rsimp (RALTS (rsa @ a # rsb)) = rsimp (RALTS (rsa @ rsb))"
Chengsong
parents: 456
diff changeset
    67
  using simp_removes_duplicate2 by auto
Chengsong
parents: 456
diff changeset
    68
Chengsong
parents: 456
diff changeset
    69
lemma distinct_removes_middle4:
Chengsong
parents: 456
diff changeset
    70
  shows "a \<in> set rsa \<Longrightarrow> rdistinct (rsa @ [a] @ rsb) rset = rdistinct (rsa @ rsb) rset"
Chengsong
parents: 456
diff changeset
    71
  using distinct_removes_middle(1) by fastforce
Chengsong
parents: 456
diff changeset
    72
Chengsong
parents: 456
diff changeset
    73
lemma distinct_removes_middle_list:
Chengsong
parents: 456
diff changeset
    74
  shows "\<forall>a \<in> set x. a \<in> set rsa \<Longrightarrow> rdistinct (rsa @ x @ rsb) rset = rdistinct (rsa @ rsb) rset"
Chengsong
parents: 456
diff changeset
    75
  apply(induct x)
Chengsong
parents: 456
diff changeset
    76
   apply simp
Chengsong
parents: 456
diff changeset
    77
  by (simp add: distinct_removes_middle3)
Chengsong
parents: 456
diff changeset
    78
Chengsong
parents: 456
diff changeset
    79
Chengsong
parents: 456
diff changeset
    80
lemma distinct_removes_duplicate_flts2:
Chengsong
parents: 456
diff changeset
    81
  shows " a \<in> set rsa
Chengsong
parents: 456
diff changeset
    82
       \<Longrightarrow> rdistinct (rflts (rsa @ [a] @ rsb)) {} =
Chengsong
parents: 456
diff changeset
    83
           rdistinct (rflts (rsa @ rsb)) {}"
Chengsong
parents: 456
diff changeset
    84
  apply(induct a arbitrary: rsb)
Chengsong
parents: 456
diff changeset
    85
  using flts_middle01 apply presburger
Chengsong
parents: 456
diff changeset
    86
      apply(subgoal_tac "rflts (rsa @ [RONE] @ rsb) = rflts rsa @ [RONE] @ rflts rsb")
Chengsong
parents: 456
diff changeset
    87
  prefer 2
Chengsong
parents: 456
diff changeset
    88
  using flts_append1 apply blast
Chengsong
parents: 456
diff changeset
    89
      apply simp
Chengsong
parents: 456
diff changeset
    90
      apply(subgoal_tac "RONE \<in> set (rflts rsa)")
Chengsong
parents: 456
diff changeset
    91
  prefer 2
Chengsong
parents: 456
diff changeset
    92
  using rflts_def_idiot2 apply blast
Chengsong
parents: 456
diff changeset
    93
      apply(subst distinct_removes_middle3)
Chengsong
parents: 456
diff changeset
    94
       apply simp
Chengsong
parents: 456
diff changeset
    95
  using flts_append apply presburger
Chengsong
parents: 456
diff changeset
    96
     apply simp
Chengsong
parents: 456
diff changeset
    97
  apply (metis distinct_removes_middle3 flts_append in_set_conv_decomp rflts.simps(5))
Chengsong
parents: 456
diff changeset
    98
  apply (metis distinct_removes_middle(1) flts_append flts_append1 rflts_def_idiot2 rrexp.distinct(25) rrexp.distinct(5))
Chengsong
parents: 456
diff changeset
    99
   apply(subgoal_tac "rflts (rsa @ [RALTS x] @ rsb) = rflts rsa @ x @ rflts rsb")
Chengsong
parents: 456
diff changeset
   100
    prefer 2
Chengsong
parents: 456
diff changeset
   101
  apply (simp add: flts_append)
Chengsong
parents: 456
diff changeset
   102
   apply (simp only:)
Chengsong
parents: 456
diff changeset
   103
Chengsong
parents: 456
diff changeset
   104
   apply(subgoal_tac "\<forall>r1 \<in> set x. r1 \<in> set (rflts rsa)")
Chengsong
parents: 456
diff changeset
   105
    prefer 2
Chengsong
parents: 456
diff changeset
   106
  using spilled_alts_contained apply blast
Chengsong
parents: 456
diff changeset
   107
  apply(subst flts_append)
Chengsong
parents: 456
diff changeset
   108
  using distinct_removes_middle_list apply blast
Chengsong
parents: 456
diff changeset
   109
  using distinct_removes_middle2 flts_append rflts_def_idiot2 by fastforce
Chengsong
parents: 456
diff changeset
   110
Chengsong
parents: 456
diff changeset
   111
Chengsong
parents: 456
diff changeset
   112
lemma simp_removes_duplicate:
Chengsong
parents: 456
diff changeset
   113
  shows "a \<in> set rsa \<Longrightarrow> rsimp (rsimp_ALTs (rsa @ a # rs)) =  rsimp (rsimp_ALTs (rsa @ rs))"
Chengsong
parents: 456
diff changeset
   114
  apply(subgoal_tac "rsimp (rsimp_ALTs (rsa @ a # rs)) = rsimp (RALTS (rsa @ a # rs))")
Chengsong
parents: 456
diff changeset
   115
   prefer 2
Chengsong
parents: 456
diff changeset
   116
  apply (smt (verit, best) Cons_eq_append_conv append_is_Nil_conv empty_set equals0D list.distinct(1) rsimp_ALTs.elims)
Chengsong
parents: 456
diff changeset
   117
  apply(simp only:)
Chengsong
parents: 456
diff changeset
   118
  apply simp
Chengsong
parents: 456
diff changeset
   119
  apply(subgoal_tac "(rdistinct (rflts (map rsimp rsa @ rsimp a # map rsimp rs)) {}) = (rdistinct (rflts (map rsimp rsa @  map rsimp rs)) {})")
Chengsong
parents: 456
diff changeset
   120
   apply(simp only:)
Chengsong
parents: 456
diff changeset
   121
  prefer 2
Chengsong
parents: 456
diff changeset
   122
   apply(subgoal_tac "rsimp a \<in> set (map rsimp rsa)")
Chengsong
parents: 456
diff changeset
   123
    prefer 2
Chengsong
parents: 456
diff changeset
   124
  apply simp
Chengsong
parents: 456
diff changeset
   125
  using distinct_removes_duplicate_flts2 apply force
Chengsong
parents: 456
diff changeset
   126
  apply(case_tac rsa)
Chengsong
parents: 456
diff changeset
   127
   apply simp
Chengsong
parents: 456
diff changeset
   128
  apply(case_tac rs)
Chengsong
parents: 456
diff changeset
   129
   apply simp
Chengsong
parents: 456
diff changeset
   130
   apply(case_tac list)
Chengsong
parents: 456
diff changeset
   131
    apply simp
Chengsong
parents: 456
diff changeset
   132
  using idem_after_simp1 apply presburger
Chengsong
parents: 456
diff changeset
   133
   apply simp+
Chengsong
parents: 456
diff changeset
   134
  apply(subgoal_tac "rsimp_ALTs (aa # list @ aaa # lista) = RALTS (aa # list @ aaa # lista)")
Chengsong
parents: 456
diff changeset
   135
   apply simp
Chengsong
parents: 456
diff changeset
   136
  using rsimpalts_conscons by presburger
467
Chengsong
parents: 465
diff changeset
   137
Chengsong
parents: 465
diff changeset
   138
Chengsong
parents: 465
diff changeset
   139
Chengsong
parents: 465
diff changeset
   140
468
a0f27e21b42c all texrelated
Chengsong
parents: 467
diff changeset
   141
467
Chengsong
parents: 465
diff changeset
   142
Chengsong
parents: 465
diff changeset
   143
471
Chengsong
parents: 469
diff changeset
   144
Chengsong
parents: 469
diff changeset
   145
inductive frewrite:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ \<leadsto>f _" [10, 10] 10)
Chengsong
parents: 469
diff changeset
   146
  where
Chengsong
parents: 469
diff changeset
   147
  "(RZERO # rs) \<leadsto>f rs"
Chengsong
parents: 469
diff changeset
   148
| "((RALTS rs) # rsa) \<leadsto>f (rs @ rsa)"
Chengsong
parents: 469
diff changeset
   149
| "rs1 \<leadsto>f rs2 \<Longrightarrow> (r # rs1) \<leadsto>f (r # rs2)"
Chengsong
parents: 469
diff changeset
   150
Chengsong
parents: 469
diff changeset
   151
Chengsong
parents: 469
diff changeset
   152
inductive 
Chengsong
parents: 469
diff changeset
   153
  frewrites:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ \<leadsto>f* _" [10, 10] 10)
Chengsong
parents: 469
diff changeset
   154
where 
473
Chengsong
parents: 472
diff changeset
   155
  [intro, simp]:"rs \<leadsto>f* rs"
Chengsong
parents: 472
diff changeset
   156
| [intro]: "\<lbrakk>rs1 \<leadsto>f* rs2; rs2 \<leadsto>f rs3\<rbrakk> \<Longrightarrow> rs1 \<leadsto>f* rs3"
Chengsong
parents: 472
diff changeset
   157
Chengsong
parents: 472
diff changeset
   158
inductive grewrite:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ \<leadsto>g _" [10, 10] 10)
Chengsong
parents: 472
diff changeset
   159
  where
Chengsong
parents: 472
diff changeset
   160
  "(RZERO # rs) \<leadsto>g rs"
Chengsong
parents: 472
diff changeset
   161
| "((RALTS rs) # rsa) \<leadsto>g (rs @ rsa)"
Chengsong
parents: 472
diff changeset
   162
| "rs1 \<leadsto>g rs2 \<Longrightarrow> (r # rs1) \<leadsto>g (r # rs2)"
Chengsong
parents: 472
diff changeset
   163
| "rsa @ [a] @ rsb @ [a] @ rsc \<leadsto>g rsa @ [a] @ rsb @ rsc" 
Chengsong
parents: 472
diff changeset
   164
478
Chengsong
parents: 476
diff changeset
   165
lemma grewrite_variant1:
Chengsong
parents: 476
diff changeset
   166
  shows "a \<in> set rs1 \<Longrightarrow> rs1 @ a # rs \<leadsto>g rs1 @ rs"
Chengsong
parents: 476
diff changeset
   167
  apply (metis append.assoc append_Cons append_Nil grewrite.intros(4) split_list_first)
Chengsong
parents: 476
diff changeset
   168
  done
Chengsong
parents: 476
diff changeset
   169
473
Chengsong
parents: 472
diff changeset
   170
Chengsong
parents: 472
diff changeset
   171
inductive 
Chengsong
parents: 472
diff changeset
   172
  grewrites:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ \<leadsto>g* _" [10, 10] 10)
Chengsong
parents: 472
diff changeset
   173
where 
Chengsong
parents: 472
diff changeset
   174
  [intro, simp]:"rs \<leadsto>g* rs"
Chengsong
parents: 472
diff changeset
   175
| [intro]: "\<lbrakk>rs1 \<leadsto>g* rs2; rs2 \<leadsto>g rs3\<rbrakk> \<Longrightarrow> rs1 \<leadsto>g* rs3"
478
Chengsong
parents: 476
diff changeset
   176
Chengsong
parents: 476
diff changeset
   177
Chengsong
parents: 476
diff changeset
   178
473
Chengsong
parents: 472
diff changeset
   179
(*
Chengsong
parents: 472
diff changeset
   180
inductive 
Chengsong
parents: 472
diff changeset
   181
  frewrites2:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ <\<leadsto>f* _" [10, 10] 10)
Chengsong
parents: 472
diff changeset
   182
where 
Chengsong
parents: 472
diff changeset
   183
 [intro]: "\<lbrakk>rs1 \<leadsto>f* rs2; rs2 \<leadsto>f* rs1\<rbrakk> \<Longrightarrow> rs1 <\<leadsto>f* rs2"
Chengsong
parents: 472
diff changeset
   184
*)
471
Chengsong
parents: 469
diff changeset
   185
Chengsong
parents: 469
diff changeset
   186
lemma fr_in_rstar : "r1 \<leadsto>f r2 \<Longrightarrow> r1 \<leadsto>f* r2"
Chengsong
parents: 469
diff changeset
   187
  using frewrites.intros(1) frewrites.intros(2) by blast
Chengsong
parents: 469
diff changeset
   188
 
Chengsong
parents: 469
diff changeset
   189
lemma freal_trans[trans]: 
Chengsong
parents: 469
diff changeset
   190
  assumes a1: "r1 \<leadsto>f* r2"  and a2: "r2 \<leadsto>f* r3"
Chengsong
parents: 469
diff changeset
   191
  shows "r1 \<leadsto>f* r3"
Chengsong
parents: 469
diff changeset
   192
  using a2 a1
Chengsong
parents: 469
diff changeset
   193
  apply(induct r2 r3 arbitrary: r1 rule: frewrites.induct) 
Chengsong
parents: 469
diff changeset
   194
  apply(auto)
Chengsong
parents: 469
diff changeset
   195
  done
Chengsong
parents: 469
diff changeset
   196
Chengsong
parents: 469
diff changeset
   197
Chengsong
parents: 469
diff changeset
   198
lemma  many_steps_later: "\<lbrakk>r1 \<leadsto>f r2; r2 \<leadsto>f* r3 \<rbrakk> \<Longrightarrow> r1 \<leadsto>f* r3"
Chengsong
parents: 469
diff changeset
   199
  by (meson fr_in_rstar freal_trans)
Chengsong
parents: 469
diff changeset
   200
Chengsong
parents: 469
diff changeset
   201
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   202
lemma gr_in_rstar : "r1 \<leadsto>g r2 \<Longrightarrow> r1 \<leadsto>g* r2"
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   203
  using grewrites.intros(1) grewrites.intros(2) by blast
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   204
 
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   205
lemma greal_trans[trans]: 
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   206
  assumes a1: "r1 \<leadsto>g* r2"  and a2: "r2 \<leadsto>g* r3"
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   207
  shows "r1 \<leadsto>g* r3"
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   208
  using a2 a1
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   209
  apply(induct r2 r3 arbitrary: r1 rule: grewrites.induct) 
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   210
  apply(auto)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   211
  done
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   212
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   213
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   214
lemma  gmany_steps_later: "\<lbrakk>r1 \<leadsto>g r2; r2 \<leadsto>g* r3 \<rbrakk> \<Longrightarrow> r1 \<leadsto>g* r3"
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   215
  by (meson gr_in_rstar greal_trans)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   216
478
Chengsong
parents: 476
diff changeset
   217
lemma gstar_rdistinct_general:
Chengsong
parents: 476
diff changeset
   218
  shows "rs1 @  rs \<leadsto>g* rs1 @ (rdistinct rs (set rs1))"
Chengsong
parents: 476
diff changeset
   219
  apply(induct rs arbitrary: rs1)
Chengsong
parents: 476
diff changeset
   220
   apply simp
Chengsong
parents: 476
diff changeset
   221
  apply(case_tac " a \<in> set rs1")
Chengsong
parents: 476
diff changeset
   222
   apply simp
Chengsong
parents: 476
diff changeset
   223
  apply(subgoal_tac "rs1 @ a # rs \<leadsto>g rs1 @ rs")
Chengsong
parents: 476
diff changeset
   224
  using gmany_steps_later apply auto[1]
Chengsong
parents: 476
diff changeset
   225
  apply (metis append.assoc append_Cons append_Nil grewrite.intros(4) split_list_first)
Chengsong
parents: 476
diff changeset
   226
  apply simp
Chengsong
parents: 476
diff changeset
   227
  apply(drule_tac x = "rs1 @ [a]" in meta_spec)
Chengsong
parents: 476
diff changeset
   228
  by simp
Chengsong
parents: 476
diff changeset
   229
Chengsong
parents: 476
diff changeset
   230
Chengsong
parents: 476
diff changeset
   231
lemma gstar_rdistinct:
Chengsong
parents: 476
diff changeset
   232
  shows "rs \<leadsto>g* rdistinct rs {}"
Chengsong
parents: 476
diff changeset
   233
  apply(induct rs)
Chengsong
parents: 476
diff changeset
   234
   apply simp
Chengsong
parents: 476
diff changeset
   235
  by (metis append.left_neutral empty_set gstar_rdistinct_general)
Chengsong
parents: 476
diff changeset
   236
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   237
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   238
471
Chengsong
parents: 469
diff changeset
   239
lemma frewrite_append:
Chengsong
parents: 469
diff changeset
   240
  shows "\<lbrakk> rsa \<leadsto>f rsb \<rbrakk> \<Longrightarrow> rs @ rsa \<leadsto>f rs @ rsb"
Chengsong
parents: 469
diff changeset
   241
  apply(induct rs)
472
Chengsong
parents: 471
diff changeset
   242
   apply simp+
Chengsong
parents: 471
diff changeset
   243
  using frewrite.intros(3) by blast
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   244
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   245
lemma grewrite_append:
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   246
  shows "\<lbrakk> rsa \<leadsto>g rsb \<rbrakk> \<Longrightarrow> rs @ rsa \<leadsto>g rs @ rsb"
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   247
  apply(induct rs)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   248
   apply simp+
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   249
  using grewrite.intros(3) by blast
472
Chengsong
parents: 471
diff changeset
   250
  
471
Chengsong
parents: 469
diff changeset
   251
Chengsong
parents: 469
diff changeset
   252
Chengsong
parents: 469
diff changeset
   253
lemma frewrites_cons:
Chengsong
parents: 469
diff changeset
   254
  shows "\<lbrakk> rsa \<leadsto>f* rsb \<rbrakk> \<Longrightarrow> r # rsa \<leadsto>f* r # rsb"
472
Chengsong
parents: 471
diff changeset
   255
  apply(induct rsa rsb rule: frewrites.induct)
Chengsong
parents: 471
diff changeset
   256
   apply simp
Chengsong
parents: 471
diff changeset
   257
  using frewrite.intros(3) by blast
471
Chengsong
parents: 469
diff changeset
   258
Chengsong
parents: 469
diff changeset
   259
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   260
lemma grewrites_cons:
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   261
  shows "\<lbrakk> rsa \<leadsto>g* rsb \<rbrakk> \<Longrightarrow> r # rsa \<leadsto>g* r # rsb"
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   262
  apply(induct rsa rsb rule: grewrites.induct)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   263
   apply simp
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   264
  using grewrite.intros(3) by blast
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   265
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   266
471
Chengsong
parents: 469
diff changeset
   267
lemma frewrites_append:
Chengsong
parents: 469
diff changeset
   268
  shows " \<lbrakk>rsa \<leadsto>f* rsb\<rbrakk> \<Longrightarrow> (rs @ rsa) \<leadsto>f* (rs @ rsb)"
472
Chengsong
parents: 471
diff changeset
   269
  apply(induct rs)
471
Chengsong
parents: 469
diff changeset
   270
   apply simp
472
Chengsong
parents: 471
diff changeset
   271
  by (simp add: frewrites_cons)
471
Chengsong
parents: 469
diff changeset
   272
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   273
lemma grewrites_append:
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   274
  shows " \<lbrakk>rsa \<leadsto>g* rsb\<rbrakk> \<Longrightarrow> (rs @ rsa) \<leadsto>g* (rs @ rsb)"
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   275
  apply(induct rs)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   276
   apply simp
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   277
  by (simp add: grewrites_cons)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   278
471
Chengsong
parents: 469
diff changeset
   279
Chengsong
parents: 469
diff changeset
   280
Chengsong
parents: 469
diff changeset
   281
lemma frewrites_concat:
Chengsong
parents: 469
diff changeset
   282
  shows "\<lbrakk>rs1 \<leadsto>f rs2; rsa \<leadsto>f* rsb \<rbrakk> \<Longrightarrow> (rs1 @ rsa) \<leadsto>f* (rs2 @ rsb)"
Chengsong
parents: 469
diff changeset
   283
  apply(induct rs1 rs2 rule: frewrite.induct)
Chengsong
parents: 469
diff changeset
   284
    apply(simp)
Chengsong
parents: 469
diff changeset
   285
  apply(subgoal_tac "(RZERO # rs @ rsa) \<leadsto>f (rs @ rsa)")
Chengsong
parents: 469
diff changeset
   286
  prefer 2
Chengsong
parents: 469
diff changeset
   287
  using frewrite.intros(1) apply blast
Chengsong
parents: 469
diff changeset
   288
    apply(subgoal_tac "(rs @ rsa) \<leadsto>f* (rs @ rsb)")
Chengsong
parents: 469
diff changeset
   289
  using many_steps_later apply blast
472
Chengsong
parents: 471
diff changeset
   290
  apply (simp add: frewrites_append)
Chengsong
parents: 471
diff changeset
   291
  apply (metis append.assoc append_Cons frewrite.intros(2) frewrites_append many_steps_later)
Chengsong
parents: 471
diff changeset
   292
  using frewrites_cons by auto
471
Chengsong
parents: 469
diff changeset
   293
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   294
lemma grewrites_concat:
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   295
  shows "\<lbrakk>rs1 \<leadsto>g rs2; rsa \<leadsto>g* rsb \<rbrakk> \<Longrightarrow> (rs1 @ rsa) \<leadsto>g* (rs2 @ rsb)"
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   296
  apply(induct rs1 rs2 rule: grewrite.induct)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   297
    apply(simp)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   298
  apply(subgoal_tac "(RZERO # rs @ rsa) \<leadsto>g (rs @ rsa)")
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   299
  prefer 2
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   300
  using grewrite.intros(1) apply blast
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   301
    apply(subgoal_tac "(rs @ rsa) \<leadsto>g* (rs @ rsb)")
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   302
  using gmany_steps_later apply blast
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   303
  apply (simp add: grewrites_append)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   304
  apply (metis append.assoc append_Cons grewrite.intros(2) grewrites_append gmany_steps_later)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   305
  using grewrites_cons apply auto
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   306
  apply(subgoal_tac "rsaa @ a # rsba @ a # rsc @ rsa \<leadsto>g* rsaa @ a # rsba @ a # rsc @ rsb")
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   307
  using grewrite.intros(4) grewrites.intros(2) apply force
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   308
  using grewrites_append by auto
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   309
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   310
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   311
lemma grewritess_concat:
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   312
  shows "\<lbrakk>rsa \<leadsto>g* rsb; rsc \<leadsto>g* rsd \<rbrakk> \<Longrightarrow> (rsa @ rsc) \<leadsto>g* (rsb @ rsd)"
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   313
  apply(induct rsa rsb rule: grewrites.induct)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   314
   apply(case_tac rs)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   315
    apply simp
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   316
  using grewrites_append apply blast   
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   317
  by (meson greal_trans grewrites.simps grewrites_concat)
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   318
476
Chengsong
parents: 475
diff changeset
   319
fun alt_set:: "rrexp \<Rightarrow> rrexp set"
Chengsong
parents: 475
diff changeset
   320
  where
Chengsong
parents: 475
diff changeset
   321
  "alt_set (RALTS rs) = set rs \<union> \<Union> (alt_set ` (set rs))"
Chengsong
parents: 475
diff changeset
   322
| "alt_set r = {r}"
Chengsong
parents: 475
diff changeset
   323
Chengsong
parents: 475
diff changeset
   324
lemma alt_set_has_all:
Chengsong
parents: 475
diff changeset
   325
  shows "RALTS rs \<in> alt_set rx \<Longrightarrow> set rs \<subseteq> alt_set rx"
Chengsong
parents: 475
diff changeset
   326
  apply(induct rx arbitrary: rs)
Chengsong
parents: 475
diff changeset
   327
       apply simp_all
Chengsong
parents: 475
diff changeset
   328
  apply(rename_tac rSS rss)
Chengsong
parents: 475
diff changeset
   329
  using in_mono by fastforce
Chengsong
parents: 475
diff changeset
   330
Chengsong
parents: 475
diff changeset
   331
Chengsong
parents: 475
diff changeset
   332
Chengsong
parents: 475
diff changeset
   333
Chengsong
parents: 475
diff changeset
   334
lemma grewrite_equal_rsimp:
Chengsong
parents: 475
diff changeset
   335
  shows "\<lbrakk>rs1 \<leadsto>g rs2; rsimp_ALTs (rdistinct (rflts (map rsimp rs1)) (rset \<union> \<Union>(alt_set ` rset))) =
Chengsong
parents: 475
diff changeset
   336
          rsimp_ALTs (rdistinct (rflts (map rsimp rs2)) (rset \<union> \<Union>(alt_set ` rset)))\<rbrakk>
Chengsong
parents: 475
diff changeset
   337
       \<Longrightarrow> rsimp_ALTs (rdistinct (rflts (rsimp r # map rsimp rs1))  (rset \<union> \<Union>(alt_set ` rset))) =
Chengsong
parents: 475
diff changeset
   338
           rsimp_ALTs (rdistinct (rflts (rsimp r # map rsimp rs2))  (rset \<union> \<Union>(alt_set ` rset)))"
Chengsong
parents: 475
diff changeset
   339
  apply(induct rs1 rs2 arbitrary:rset rule: grewrite.induct)
Chengsong
parents: 475
diff changeset
   340
     apply simp
Chengsong
parents: 475
diff changeset
   341
  apply (metis append_Cons append_Nil flts_middle0)
Chengsong
parents: 475
diff changeset
   342
    apply(case_tac "rsimp r \<in> rset")
Chengsong
parents: 475
diff changeset
   343
     apply simp
Chengsong
parents: 475
diff changeset
   344
  oops
Chengsong
parents: 475
diff changeset
   345
478
Chengsong
parents: 476
diff changeset
   346
lemma grewrite_cases_middle:
Chengsong
parents: 476
diff changeset
   347
  shows "rs1 \<leadsto>g rs2 \<Longrightarrow> 
Chengsong
parents: 476
diff changeset
   348
(\<exists>rsa rsb rsc. rs1 =  (rsa @ [RALTS rsb] @ rsc) \<and> rs2 = (rsa @ rsb @ rsc)) \<or>
Chengsong
parents: 476
diff changeset
   349
(\<exists>rsa rsc. rs1 = rsa @ [RZERO] @ rsc \<and> rs2 = rsa @ rsc) \<or>
Chengsong
parents: 476
diff changeset
   350
(\<exists>rsa rsb rsc a. rs1 = rsa @ [a] @ rsb @ [a] @ rsc \<and> rs2 = rsa @ [a] @ rsb @ rsc)"
Chengsong
parents: 476
diff changeset
   351
  apply( induct rs1 rs2 rule: grewrite.induct)
Chengsong
parents: 476
diff changeset
   352
     apply simp
Chengsong
parents: 476
diff changeset
   353
  apply blast
Chengsong
parents: 476
diff changeset
   354
  apply (metis append_Cons append_Nil)
Chengsong
parents: 476
diff changeset
   355
  apply (metis append_Cons)
Chengsong
parents: 476
diff changeset
   356
  by blast
476
Chengsong
parents: 475
diff changeset
   357
Chengsong
parents: 475
diff changeset
   358
Chengsong
parents: 475
diff changeset
   359
lemma grewrite_equal_rsimp:
Chengsong
parents: 475
diff changeset
   360
  shows "rs1 \<leadsto>g rs2 \<Longrightarrow> rsimp (RALTS rs1) = rsimp (RALTS rs2)"
478
Chengsong
parents: 476
diff changeset
   361
  apply(frule grewrite_cases_middle)
Chengsong
parents: 476
diff changeset
   362
  apply(case_tac "(\<exists>rsa rsb rsc. rs1 = rsa @ [RALTS rsb] @ rsc \<and> rs2 = rsa @ rsb @ rsc)")  
Chengsong
parents: 476
diff changeset
   363
  using simp_flatten3 apply auto[1]
Chengsong
parents: 476
diff changeset
   364
  apply(case_tac "(\<exists>rsa rsc. rs1 = rsa @ [RZERO] @ rsc \<and> rs2 = rsa @ rsc)")
Chengsong
parents: 476
diff changeset
   365
  apply (metis (mono_tags, opaque_lifting) append_Cons append_Nil list.set_intros(1) list.simps(9) rflts.simps(2) rsimp.simps(2) rsimp.simps(3) simp_removes_duplicate3)
Chengsong
parents: 476
diff changeset
   366
  by (smt (verit) append.assoc append_Cons append_Nil in_set_conv_decomp simp_removes_duplicate3)
476
Chengsong
parents: 475
diff changeset
   367
Chengsong
parents: 475
diff changeset
   368
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   369
lemma grewrites_equal_rsimp:
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   370
  shows "rs1 \<leadsto>g* rs2 \<Longrightarrow> rsimp (RALTS rs1) = rsimp (RALTS rs2)"
476
Chengsong
parents: 475
diff changeset
   371
  apply (induct rs1 rs2 rule: grewrites.induct)
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   372
  apply simp
476
Chengsong
parents: 475
diff changeset
   373
  using grewrite_equal_rsimp by presburger
Chengsong
parents: 475
diff changeset
   374
  
Chengsong
parents: 475
diff changeset
   375
Chengsong
parents: 475
diff changeset
   376
Chengsong
parents: 475
diff changeset
   377
Chengsong
parents: 475
diff changeset
   378
lemma grewrites_equal_simp_2:
Chengsong
parents: 475
diff changeset
   379
  shows "rsimp (RALTS rs1) = rsimp (RALTS rs2) \<Longrightarrow> rs1 \<leadsto>g* rs2"
478
Chengsong
parents: 476
diff changeset
   380
  oops
Chengsong
parents: 476
diff changeset
   381
Chengsong
parents: 476
diff changeset
   382
Chengsong
parents: 476
diff changeset
   383
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   384
476
Chengsong
parents: 475
diff changeset
   385
lemma grewrites_last:
Chengsong
parents: 475
diff changeset
   386
  shows "r # [RALTS rs] \<leadsto>g*  r # rs"
Chengsong
parents: 475
diff changeset
   387
  by (metis gr_in_rstar grewrite.intros(2) grewrite.intros(3) self_append_conv)
Chengsong
parents: 475
diff changeset
   388
Chengsong
parents: 475
diff changeset
   389
lemma simp_flatten2:
Chengsong
parents: 475
diff changeset
   390
  shows "rsimp (RALTS (r # [RALTS rs])) = rsimp (RALTS (r # rs))"
Chengsong
parents: 475
diff changeset
   391
  using grewrites_equal_rsimp grewrites_last by blast
475
10fd25fba2ba identities
Chengsong
parents: 473
diff changeset
   392
473
Chengsong
parents: 472
diff changeset
   393
lemma frewrites_middle:
Chengsong
parents: 472
diff changeset
   394
  shows "rs1 \<leadsto>f* rs2 \<Longrightarrow> r # (RALTS rs # rs1) \<leadsto>f* r # (rs @ rs1)"
Chengsong
parents: 472
diff changeset
   395
  by (simp add: fr_in_rstar frewrite.intros(2) frewrite.intros(3))
471
Chengsong
parents: 469
diff changeset
   396
473
Chengsong
parents: 472
diff changeset
   397
lemma frewrites_alt:
Chengsong
parents: 472
diff changeset
   398
  shows "rs1 \<leadsto>f* rs2 \<Longrightarrow> (RALT r1 r2) # rs1 \<leadsto>f* r1 # r2 # rs2"  
Chengsong
parents: 472
diff changeset
   399
  by (metis Cons_eq_appendI append_self_conv2 frewrite.intros(2) frewrites_cons many_steps_later)
471
Chengsong
parents: 469
diff changeset
   400
Chengsong
parents: 469
diff changeset
   401
lemma early_late_der_frewrites:
Chengsong
parents: 469
diff changeset
   402
  shows "map (rder x) (rflts rs) \<leadsto>f* rflts (map (rder x) rs)"
Chengsong
parents: 469
diff changeset
   403
  apply(induct rs)
Chengsong
parents: 469
diff changeset
   404
   apply simp
Chengsong
parents: 469
diff changeset
   405
  apply(case_tac a)
Chengsong
parents: 469
diff changeset
   406
       apply simp+
Chengsong
parents: 469
diff changeset
   407
  using frewrite.intros(1) many_steps_later apply blast
Chengsong
parents: 469
diff changeset
   408
     apply(case_tac "x = x3")
472
Chengsong
parents: 471
diff changeset
   409
      apply simp
Chengsong
parents: 471
diff changeset
   410
  using frewrites_cons apply presburger
Chengsong
parents: 471
diff changeset
   411
  using frewrite.intros(1) many_steps_later apply fastforce
Chengsong
parents: 471
diff changeset
   412
  apply(case_tac "rnullable x41")
473
Chengsong
parents: 472
diff changeset
   413
     apply simp+
Chengsong
parents: 472
diff changeset
   414
     apply (simp add: frewrites_alt)
Chengsong
parents: 472
diff changeset
   415
  apply (simp add: frewrites_cons)
Chengsong
parents: 472
diff changeset
   416
   apply (simp add: frewrites_append)
Chengsong
parents: 472
diff changeset
   417
  by (simp add: frewrites_cons)
471
Chengsong
parents: 469
diff changeset
   418
Chengsong
parents: 469
diff changeset
   419
Chengsong
parents: 469
diff changeset
   420
Chengsong
parents: 469
diff changeset
   421
473
Chengsong
parents: 472
diff changeset
   422
lemma with_wo0_distinct:
Chengsong
parents: 472
diff changeset
   423
  shows "rdistinct rs rset \<leadsto>f* rdistinct rs (insert RZERO rset)"
Chengsong
parents: 472
diff changeset
   424
  apply(induct rs arbitrary: rset)
Chengsong
parents: 472
diff changeset
   425
   apply simp
Chengsong
parents: 472
diff changeset
   426
  apply(case_tac a)
Chengsong
parents: 472
diff changeset
   427
  apply(case_tac "RZERO \<in> rset")
Chengsong
parents: 472
diff changeset
   428
        apply simp+
Chengsong
parents: 472
diff changeset
   429
  using fr_in_rstar frewrite.intros(1) apply presburger
Chengsong
parents: 472
diff changeset
   430
  apply (case_tac "RONE \<in> rset")
Chengsong
parents: 472
diff changeset
   431
       apply simp+
Chengsong
parents: 472
diff changeset
   432
  using frewrites_cons apply presburger
Chengsong
parents: 472
diff changeset
   433
     apply(case_tac "a \<in> rset")
Chengsong
parents: 472
diff changeset
   434
  apply simp
Chengsong
parents: 472
diff changeset
   435
  apply (simp add: frewrites_cons)
Chengsong
parents: 472
diff changeset
   436
     apply(case_tac "a \<in> rset")
Chengsong
parents: 472
diff changeset
   437
  apply simp
Chengsong
parents: 472
diff changeset
   438
  apply (simp add: frewrites_cons)
Chengsong
parents: 472
diff changeset
   439
       apply(case_tac "a \<in> rset")
Chengsong
parents: 472
diff changeset
   440
  apply simp
Chengsong
parents: 472
diff changeset
   441
  apply (simp add: frewrites_cons)
Chengsong
parents: 472
diff changeset
   442
       apply(case_tac "a \<in> rset")
Chengsong
parents: 472
diff changeset
   443
  apply simp
Chengsong
parents: 472
diff changeset
   444
  apply (simp add: frewrites_cons)
Chengsong
parents: 472
diff changeset
   445
  done
Chengsong
parents: 472
diff changeset
   446
476
Chengsong
parents: 475
diff changeset
   447
(*Interesting lemma: not obvious but easily proven by sledgehammer*)
473
Chengsong
parents: 472
diff changeset
   448
476
Chengsong
parents: 475
diff changeset
   449
  
Chengsong
parents: 475
diff changeset
   450
Chengsong
parents: 475
diff changeset
   451
Chengsong
parents: 475
diff changeset
   452
(*lemma induction last rule not go through
Chengsong
parents: 475
diff changeset
   453
 example:
Chengsong
parents: 475
diff changeset
   454
r #
Chengsong
parents: 475
diff changeset
   455
           rdistinct rs1
Chengsong
parents: 475
diff changeset
   456
            (insert RZERO
Chengsong
parents: 475
diff changeset
   457
              (insert r
Chengsong
parents: 475
diff changeset
   458
                (rset \<union>
Chengsong
parents: 475
diff changeset
   459
                 \<Union> (alt_set `
Chengsong
parents: 475
diff changeset
   460
                     rset)))) \<leadsto>g* r #
Chengsong
parents: 475
diff changeset
   461
                                   rdistinct rs2
Chengsong
parents: 475
diff changeset
   462
                                    (insert RZERO (insert r (rset \<union> \<Union> (alt_set ` rset))))
Chengsong
parents: 475
diff changeset
   463
 rs2 = [+rs] rs3 = rs,
Chengsong
parents: 475
diff changeset
   464
r = +rs
Chengsong
parents: 475
diff changeset
   465
[] \<leadsto>g* rs which is wrong
Chengsong
parents: 475
diff changeset
   466
*)
Chengsong
parents: 475
diff changeset
   467
469
e5dd8cc0aa82 all electron pics removed
Chengsong
parents: 468
diff changeset
   468
e5dd8cc0aa82 all electron pics removed
Chengsong
parents: 468
diff changeset
   469
476
Chengsong
parents: 475
diff changeset
   470
Chengsong
parents: 475
diff changeset
   471
Chengsong
parents: 475
diff changeset
   472
Chengsong
parents: 475
diff changeset
   473
Chengsong
parents: 475
diff changeset
   474
lemma frewrite_simpeq:
Chengsong
parents: 475
diff changeset
   475
  shows "rs1 \<leadsto>f rs2 \<Longrightarrow> rsimp (RALTS rs1) = rsimp (RALTS rs2)"
Chengsong
parents: 475
diff changeset
   476
  apply(induct rs1 rs2 rule: frewrite.induct)
Chengsong
parents: 475
diff changeset
   477
    apply simp
Chengsong
parents: 475
diff changeset
   478
  using simp_flatten apply presburger
478
Chengsong
parents: 476
diff changeset
   479
  by (metis (no_types, opaque_lifting) grewrites_equal_rsimp grewrites_last list.simps(9) rsimp.simps(2))
Chengsong
parents: 476
diff changeset
   480
Chengsong
parents: 476
diff changeset
   481
lemma gstar0:
Chengsong
parents: 476
diff changeset
   482
  shows "rsa @ (rdistinct rs (set rsa)) \<leadsto>g* rsa @ (rdistinct rs (insert RZERO (set rsa)))"
Chengsong
parents: 476
diff changeset
   483
  apply(induct rs arbitrary: rsa)
Chengsong
parents: 476
diff changeset
   484
   apply simp
Chengsong
parents: 476
diff changeset
   485
  apply(case_tac "a = RZERO")
Chengsong
parents: 476
diff changeset
   486
   apply simp
Chengsong
parents: 476
diff changeset
   487
  
Chengsong
parents: 476
diff changeset
   488
  using gr_in_rstar grewrite.intros(1) grewrites_append apply presburger
Chengsong
parents: 476
diff changeset
   489
  apply(case_tac "a \<in> set rsa")
Chengsong
parents: 476
diff changeset
   490
   apply simp+
Chengsong
parents: 476
diff changeset
   491
  apply(drule_tac x = "rsa @ [a]" in meta_spec)
Chengsong
parents: 476
diff changeset
   492
  by simp
Chengsong
parents: 476
diff changeset
   493
  
Chengsong
parents: 476
diff changeset
   494
lemma gstar01:
Chengsong
parents: 476
diff changeset
   495
  shows "rdistinct rs {} \<leadsto>g* rdistinct rs {RZERO}"
Chengsong
parents: 476
diff changeset
   496
  by (metis empty_set gstar0 self_append_conv2)
Chengsong
parents: 476
diff changeset
   497
Chengsong
parents: 476
diff changeset
   498
Chengsong
parents: 476
diff changeset
   499
lemma grewrite_rdistinct_aux:
Chengsong
parents: 476
diff changeset
   500
  shows "rs @ rdistinct rsa rset \<leadsto>g* rs @ rdistinct rsa (rset \<union> set rs)"
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   501
  apply(induct rsa arbitrary: rs rset)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   502
   apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   503
  apply(case_tac " a \<in> rset")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   504
   apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   505
  apply(case_tac "a \<in> set rs")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   506
  apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   507
   apply (metis Un_insert_left Un_insert_right gmany_steps_later grewrite_variant1 insert_absorb)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   508
  apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   509
  apply(drule_tac x = "rs @ [a]" in meta_spec)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   510
  by (metis Un_insert_left Un_insert_right append.assoc append.right_neutral append_Cons append_Nil insert_absorb2 list.simps(15) set_append)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   511
  
478
Chengsong
parents: 476
diff changeset
   512
Chengsong
parents: 476
diff changeset
   513
lemma grewrite_rdistinct_worth1:
Chengsong
parents: 476
diff changeset
   514
  shows "(rsb @ [a]) @ rdistinct rs set1 \<leadsto>g* (rsb @ [a]) @ rdistinct rs (insert a set1)"
Chengsong
parents: 476
diff changeset
   515
  by (metis append.assoc empty_set grewrite_rdistinct_aux grewrites_append inf_sup_aci(5) insert_is_Un list.simps(15))
Chengsong
parents: 476
diff changeset
   516
Chengsong
parents: 476
diff changeset
   517
lemma grewrite_rdisitinct:
Chengsong
parents: 476
diff changeset
   518
  shows "rs @ rdistinct rsa {RALTS rs} \<leadsto>g* rs @ rdistinct rsa (insert (RALTS rs) (set rs))"
Chengsong
parents: 476
diff changeset
   519
  apply(induct rsa arbitrary: rs)
Chengsong
parents: 476
diff changeset
   520
   apply simp
Chengsong
parents: 476
diff changeset
   521
  apply(case_tac "a = RALTS rs")
Chengsong
parents: 476
diff changeset
   522
  apply simp
Chengsong
parents: 476
diff changeset
   523
  apply(case_tac "a \<in> set rs")
Chengsong
parents: 476
diff changeset
   524
   apply simp
Chengsong
parents: 476
diff changeset
   525
  apply(subgoal_tac "rs @
Chengsong
parents: 476
diff changeset
   526
           a # rdistinct rsa {RALTS rs, a} \<leadsto>g rs @ rdistinct rsa {RALTS rs, a}")
Chengsong
parents: 476
diff changeset
   527
    apply(subgoal_tac 
Chengsong
parents: 476
diff changeset
   528
"rs @ rdistinct rsa {RALTS rs, a} \<leadsto>g* rs @ rdistinct rsa (insert (RALTS rs) (set rs))")
Chengsong
parents: 476
diff changeset
   529
  using gmany_steps_later apply blast
Chengsong
parents: 476
diff changeset
   530
    apply(subgoal_tac 
Chengsong
parents: 476
diff changeset
   531
" rs @ rdistinct rsa {RALTS rs, a} \<leadsto>g* rs @ rdistinct rsa ({RALTS rs, a} \<union> set rs)")
Chengsong
parents: 476
diff changeset
   532
     apply (simp add: insert_absorb)
Chengsong
parents: 476
diff changeset
   533
  using grewrite_rdistinct_aux apply blast
Chengsong
parents: 476
diff changeset
   534
  using grewrite_variant1 apply blast
Chengsong
parents: 476
diff changeset
   535
  by (metis grewrite_rdistinct_aux insert_is_Un)
Chengsong
parents: 476
diff changeset
   536
Chengsong
parents: 476
diff changeset
   537
  
Chengsong
parents: 476
diff changeset
   538
lemma frewrite_rd_grewrites_general:
Chengsong
parents: 476
diff changeset
   539
  shows "\<lbrakk>rs1 \<leadsto>f rs2; \<And>rs. \<exists>rs3. 
Chengsong
parents: 476
diff changeset
   540
(rs @ (rdistinct rs1 (set rs)) \<leadsto>g* rs3) \<and> (rs @ (rdistinct rs2 (set rs)) \<leadsto>g* rs3)\<rbrakk>
Chengsong
parents: 476
diff changeset
   541
       \<Longrightarrow> 
Chengsong
parents: 476
diff changeset
   542
\<exists>rs3. (rs @ (r # rdistinct rs1 (set rs \<union> {r})) \<leadsto>g* rs3) \<and> (rs @ (r # rdistinct rs2 (set rs \<union> {r})) \<leadsto>g* rs3)"
Chengsong
parents: 476
diff changeset
   543
  apply(drule_tac x = "rs @ [r]" in meta_spec )
Chengsong
parents: 476
diff changeset
   544
  by simp
Chengsong
parents: 476
diff changeset
   545
Chengsong
parents: 476
diff changeset
   546
479
Chengsong
parents: 478
diff changeset
   547
lemma alts_g_can_flts:
Chengsong
parents: 478
diff changeset
   548
  shows "RALTS rs \<in> set rsb \<Longrightarrow> \<exists>rs1 rs2.( rflts rsb = rs1 @ rs @ rs2)"
Chengsong
parents: 478
diff changeset
   549
  by (metis flts_append rflts.simps(3) split_list_last)
Chengsong
parents: 478
diff changeset
   550
Chengsong
parents: 478
diff changeset
   551
lemma flts_gstar:
Chengsong
parents: 478
diff changeset
   552
  shows "rs \<leadsto>g* rflts rs"
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   553
  apply(induct rs)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   554
   apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   555
  apply(case_tac "a = RZERO")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   556
   apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   557
  using gmany_steps_later grewrite.intros(1) apply blast
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   558
  apply(case_tac "\<exists>rsa. a = RALTS rsa")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   559
   apply(erule exE)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   560
  apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   561
   apply (meson grewrite.intros(2) grewrites.simps grewrites_cons)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   562
  by (simp add: grewrites_cons rflts_def_idiot)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   563
479
Chengsong
parents: 478
diff changeset
   564
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   565
lemma wrong1:
479
Chengsong
parents: 478
diff changeset
   566
  shows "a \<in> set rs1 \<Longrightarrow> rs1 @ (rdistinct rs (insert a rset)) \<leadsto>g* rs1 @ (rdistinct rs (rset))"
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   567
  
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   568
  oops
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   569
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   570
lemma more_distinct1:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   571
  shows "       \<lbrakk>\<And>rsb rset rset2.
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   572
           rset2 \<subseteq> set rsb \<Longrightarrow> rsb @ rdistinct rs rset \<leadsto>g* rsb @ rdistinct rs (rset \<union> rset2);
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   573
        rset2 \<subseteq> set rsb; a \<notin> rset; a \<in> rset2\<rbrakk>
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   574
       \<Longrightarrow> rsb @ a # rdistinct rs (insert a rset) \<leadsto>g* rsb @ rdistinct rs (rset \<union> rset2)"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   575
  apply(subgoal_tac "rsb @ a # rdistinct rs (insert a rset) \<leadsto>g* rsb @ rdistinct rs (insert a rset)")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   576
   apply(subgoal_tac "rsb @ rdistinct rs (insert a rset) \<leadsto>g* rsb @ rdistinct rs (rset \<union> rset2)")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   577
    apply (meson greal_trans)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   578
   apply (metis Un_iff Un_insert_left insert_absorb)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   579
  by (simp add: gr_in_rstar grewrite_variant1 in_mono)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   580
  
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   581
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   582
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   583
478
Chengsong
parents: 476
diff changeset
   584
479
Chengsong
parents: 478
diff changeset
   585
lemma grewrites_in_distinct0:
Chengsong
parents: 478
diff changeset
   586
  shows "a \<in>  set rs1 \<Longrightarrow> rs1 @ (rdistinct (a # rs) rset) \<leadsto>g* rs1 @ (rdistinct rs rset)"
Chengsong
parents: 478
diff changeset
   587
  apply(case_tac "a \<in> rset")
Chengsong
parents: 478
diff changeset
   588
   apply simp
Chengsong
parents: 478
diff changeset
   589
  apply simp
Chengsong
parents: 478
diff changeset
   590
  oops
Chengsong
parents: 478
diff changeset
   591
Chengsong
parents: 478
diff changeset
   592
 
Chengsong
parents: 478
diff changeset
   593
Chengsong
parents: 478
diff changeset
   594
Chengsong
parents: 478
diff changeset
   595
  
Chengsong
parents: 478
diff changeset
   596
Chengsong
parents: 478
diff changeset
   597
478
Chengsong
parents: 476
diff changeset
   598
Chengsong
parents: 476
diff changeset
   599
Chengsong
parents: 476
diff changeset
   600
lemma frewrite_rd_grewrites_aux:
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   601
  shows     "       RALTS rs \<notin> set rsb \<Longrightarrow>
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   602
       rsb @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   603
       RALTS rs #
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   604
       rdistinct rsa
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   605
        (insert (RALTS rs)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   606
          (set rsb)) \<leadsto>g* rflts rsb @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   607
                          rdistinct rs (set rsb) @ rdistinct rsa (set rs \<union> set rsb \<union> {RALTS rs})"
478
Chengsong
parents: 476
diff changeset
   608
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   609
   apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   610
  apply(subgoal_tac "rsb @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   611
    RALTS rs #
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   612
    rdistinct rsa
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   613
     (insert (RALTS rs)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   614
       (set rsb)) \<leadsto>g* rsb @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   615
    rs @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   616
    rdistinct rsa
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   617
     (insert (RALTS rs)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   618
       (set rsb)) ")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   619
  apply(subgoal_tac " rsb @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   620
    rs @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   621
    rdistinct rsa
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   622
     (insert (RALTS rs)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   623
       (set rsb)) \<leadsto>g*
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   624
                      rsb @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   625
    rdistinct rs (set rsb) @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   626
    rdistinct rsa
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   627
     (insert (RALTS rs)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   628
       (set rsb)) ")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   629
    apply (smt (verit, ccfv_SIG) Un_insert_left flts_gstar greal_trans grewrite_rdistinct_aux grewritess_concat inf_sup_aci(5) rdistinct_concat_general rdistinct_set_equality set_append)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   630
   apply (metis append_assoc grewrites.intros(1) grewritess_concat gstar_rdistinct_general)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   631
  by (simp add: gr_in_rstar grewrite.intros(2) grewrites_append)
478
Chengsong
parents: 476
diff changeset
   632
  
Chengsong
parents: 476
diff changeset
   633
479
Chengsong
parents: 478
diff changeset
   634
478
Chengsong
parents: 476
diff changeset
   635
Chengsong
parents: 476
diff changeset
   636
lemma list_dlist_union:
Chengsong
parents: 476
diff changeset
   637
  shows "set rs \<subseteq> set rsb \<union> set (rdistinct rs (set rsb))"
Chengsong
parents: 476
diff changeset
   638
  by (metis rdistinct_concat_general rdistinct_set_equality set_append sup_ge2)
Chengsong
parents: 476
diff changeset
   639
Chengsong
parents: 476
diff changeset
   640
lemma subset_distinct_rewrite1:
Chengsong
parents: 476
diff changeset
   641
  shows "set1 \<subseteq> set rsb \<Longrightarrow> rsb @ rs \<leadsto>g* rsb @ (rdistinct rs set1)"
Chengsong
parents: 476
diff changeset
   642
  apply(induct rs arbitrary: rsb)
Chengsong
parents: 476
diff changeset
   643
   apply simp
Chengsong
parents: 476
diff changeset
   644
  apply(case_tac "a \<in> set1")
Chengsong
parents: 476
diff changeset
   645
   apply simp
Chengsong
parents: 476
diff changeset
   646
  
Chengsong
parents: 476
diff changeset
   647
  using gmany_steps_later grewrite_variant1 apply blast
Chengsong
parents: 476
diff changeset
   648
  apply simp
Chengsong
parents: 476
diff changeset
   649
  apply(drule_tac x = "rsb @ [a]" in meta_spec)
Chengsong
parents: 476
diff changeset
   650
  apply(subgoal_tac "set1 \<subseteq> set (rsb @ [a])")
Chengsong
parents: 476
diff changeset
   651
   apply (simp only:)
Chengsong
parents: 476
diff changeset
   652
  apply(subgoal_tac "(rsb @ [a]) @ rdistinct rs set1 \<leadsto>g* (rsb @ [a]) @ rdistinct rs (insert a set1)")
Chengsong
parents: 476
diff changeset
   653
    apply (metis (no_types, opaque_lifting) append.assoc append_Cons append_Nil greal_trans)
Chengsong
parents: 476
diff changeset
   654
  apply (metis append.assoc empty_set grewrite_rdistinct_aux grewrites_append inf_sup_aci(5) insert_is_Un list.simps(15))
Chengsong
parents: 476
diff changeset
   655
  by auto
Chengsong
parents: 476
diff changeset
   656
Chengsong
parents: 476
diff changeset
   657
Chengsong
parents: 476
diff changeset
   658
lemma subset_distinct_rewrite:
Chengsong
parents: 476
diff changeset
   659
  shows "set rsb' \<subseteq> set rsb \<Longrightarrow> rsb @ rs \<leadsto>g* rsb @ (rdistinct rs (set rsb'))"
Chengsong
parents: 476
diff changeset
   660
  by (simp add: subset_distinct_rewrite1)
Chengsong
parents: 476
diff changeset
   661
  
Chengsong
parents: 476
diff changeset
   662
Chengsong
parents: 476
diff changeset
   663
Chengsong
parents: 476
diff changeset
   664
lemma distinct_3list:
Chengsong
parents: 476
diff changeset
   665
  shows "rsb @ (rdistinct rs (set rsb)) @ rsa \<leadsto>g* 
Chengsong
parents: 476
diff changeset
   666
         rsb @ (rdistinct rs (set rsb)) @ (rdistinct rsa (set rs))"
Chengsong
parents: 476
diff changeset
   667
  by (metis append.assoc list_dlist_union set_append subset_distinct_rewrite)
Chengsong
parents: 476
diff changeset
   668
  
Chengsong
parents: 476
diff changeset
   669
Chengsong
parents: 476
diff changeset
   670
Chengsong
parents: 476
diff changeset
   671
Chengsong
parents: 476
diff changeset
   672
lemma grewrites_shape1:
Chengsong
parents: 476
diff changeset
   673
  shows "      RALTS rs \<notin> set rsb \<Longrightarrow>
Chengsong
parents: 476
diff changeset
   674
       rsb @
Chengsong
parents: 476
diff changeset
   675
       RALTS rs #
Chengsong
parents: 476
diff changeset
   676
       rdistinct rsa
Chengsong
parents: 476
diff changeset
   677
        (
Chengsong
parents: 476
diff changeset
   678
          (set rsb)) \<leadsto>g* rsb @
Chengsong
parents: 476
diff changeset
   679
                          rdistinct rs (set rsb) @
Chengsong
parents: 476
diff changeset
   680
                          rdistinct (rflts (rdistinct rsa ( (set rsb \<union> set rs)))) (set rs)"
Chengsong
parents: 476
diff changeset
   681
Chengsong
parents: 476
diff changeset
   682
Chengsong
parents: 476
diff changeset
   683
  apply (subgoal_tac "       rsb @
Chengsong
parents: 476
diff changeset
   684
       RALTS rs #
Chengsong
parents: 476
diff changeset
   685
       rdistinct rsa
Chengsong
parents: 476
diff changeset
   686
        (
Chengsong
parents: 476
diff changeset
   687
          (set rsb))  \<leadsto>g*        rsb @
Chengsong
parents: 476
diff changeset
   688
       rs @
Chengsong
parents: 476
diff changeset
   689
       rdistinct rsa
Chengsong
parents: 476
diff changeset
   690
        (
Chengsong
parents: 476
diff changeset
   691
          (set rsb)) ")
Chengsong
parents: 476
diff changeset
   692
   prefer 2
Chengsong
parents: 476
diff changeset
   693
  using gr_in_rstar grewrite.intros(2) grewrites_append apply presburger
Chengsong
parents: 476
diff changeset
   694
  apply(subgoal_tac "rsb @ rs @ rdistinct rsa (  (set rsb)) \<leadsto>g* rsb @ 
Chengsong
parents: 476
diff changeset
   695
(rdistinct rs (set rsb) @ rdistinct rsa (  (set rsb)))")
Chengsong
parents: 476
diff changeset
   696
  prefer 2
Chengsong
parents: 476
diff changeset
   697
  apply (metis append_assoc grewrites.intros(1) grewritess_concat gstar_rdistinct_general)
Chengsong
parents: 476
diff changeset
   698
  apply(subgoal_tac " rsb @ rdistinct rs (set rsb) @ rdistinct rsa (  (set rsb))
Chengsong
parents: 476
diff changeset
   699
\<leadsto>g*  rsb @ rdistinct rs (set rsb) @ rdistinct rsa (  (set rsb) \<union> (set rs))")
Chengsong
parents: 476
diff changeset
   700
  prefer 2
Chengsong
parents: 476
diff changeset
   701
  apply (smt (verit, best) append.assoc append_assoc boolean_algebra_cancel.sup2 grewrite_rdistinct_aux inf_sup_aci(5) insert_is_Un rdistinct_concat_general rdistinct_set_equality set_append sup.commute sup.right_idem sup_commute)
Chengsong
parents: 476
diff changeset
   702
  apply(subgoal_tac "rdistinct rsa (  (set rsb)  \<union> set rs) \<leadsto>g*
Chengsong
parents: 476
diff changeset
   703
rflts (rdistinct rsa (  (set rsb) \<union> set rs))")
Chengsong
parents: 476
diff changeset
   704
   apply(subgoal_tac "rsb @ (rdistinct rs (set rsb)) @ rflts (rdistinct rsa (  (set rsb) \<union> set rs)) \<leadsto>g* 
Chengsong
parents: 476
diff changeset
   705
rsb @ (rdistinct rs (set rsb)) @ (rdistinct (rflts (rdistinct rsa (  (set rsb) \<union> set rs))) (set rs))")
Chengsong
parents: 476
diff changeset
   706
  apply (smt (verit, ccfv_SIG) Un_insert_left greal_trans grewrites_append)
Chengsong
parents: 476
diff changeset
   707
  using distinct_3list apply presburger
Chengsong
parents: 476
diff changeset
   708
  using flts_gstar apply blast
Chengsong
parents: 476
diff changeset
   709
  done
Chengsong
parents: 476
diff changeset
   710
Chengsong
parents: 476
diff changeset
   711
lemma r_finite1:
Chengsong
parents: 476
diff changeset
   712
  shows "r = RALTS (r # rs) = False"
Chengsong
parents: 476
diff changeset
   713
  apply(induct r)
Chengsong
parents: 476
diff changeset
   714
  apply simp+
Chengsong
parents: 476
diff changeset
   715
   apply (metis list.set_intros(1))
Chengsong
parents: 476
diff changeset
   716
  by blast
Chengsong
parents: 476
diff changeset
   717
  
Chengsong
parents: 476
diff changeset
   718
Chengsong
parents: 476
diff changeset
   719
Chengsong
parents: 476
diff changeset
   720
lemma grewrite_singleton:
Chengsong
parents: 476
diff changeset
   721
  shows "[r] \<leadsto>g r # rs \<Longrightarrow> rs = []"
Chengsong
parents: 476
diff changeset
   722
  apply (induct "[r]" "r # rs" rule: grewrite.induct)
Chengsong
parents: 476
diff changeset
   723
    apply simp
Chengsong
parents: 476
diff changeset
   724
  apply (metis r_finite1)
Chengsong
parents: 476
diff changeset
   725
  using grewrite.simps apply blast
Chengsong
parents: 476
diff changeset
   726
  by simp
Chengsong
parents: 476
diff changeset
   727
Chengsong
parents: 476
diff changeset
   728
lemma impossible_grewrite1:
Chengsong
parents: 476
diff changeset
   729
  shows "\<not>( [RONE] \<leadsto>g [])"
Chengsong
parents: 476
diff changeset
   730
  using grewrite.cases by fastforce
Chengsong
parents: 476
diff changeset
   731
Chengsong
parents: 476
diff changeset
   732
Chengsong
parents: 476
diff changeset
   733
lemma impossible_grewrite2:
Chengsong
parents: 476
diff changeset
   734
  shows "\<not> ([RALTS rs] \<leadsto>g (RALTS rs) # a # rs)"
Chengsong
parents: 476
diff changeset
   735
  using grewrite_singleton by blast
Chengsong
parents: 476
diff changeset
   736
Chengsong
parents: 476
diff changeset
   737
  
Chengsong
parents: 476
diff changeset
   738
Chengsong
parents: 476
diff changeset
   739
479
Chengsong
parents: 478
diff changeset
   740
lemma wront_sublist_grewrites:
478
Chengsong
parents: 476
diff changeset
   741
  shows "rs1 @ rs2 \<leadsto>g* rs1 @ rs3 \<Longrightarrow> rs2 \<leadsto>g* rs3"
Chengsong
parents: 476
diff changeset
   742
  apply(induct rs1 arbitrary: rs2 rs3 rule: rev_induct)
Chengsong
parents: 476
diff changeset
   743
   apply simp
Chengsong
parents: 476
diff changeset
   744
  apply(drule_tac x = "[x] @ rs2" in meta_spec)
Chengsong
parents: 476
diff changeset
   745
  apply(drule_tac x = "[x] @ rs3" in meta_spec)
Chengsong
parents: 476
diff changeset
   746
  apply(simp)
Chengsong
parents: 476
diff changeset
   747
Chengsong
parents: 476
diff changeset
   748
  oops
Chengsong
parents: 476
diff changeset
   749
Chengsong
parents: 476
diff changeset
   750
Chengsong
parents: 476
diff changeset
   751
Chengsong
parents: 476
diff changeset
   752
479
Chengsong
parents: 478
diff changeset
   753
lemma concat_rdistinct_equality1:
Chengsong
parents: 478
diff changeset
   754
  shows "rdistinct (rs @ rsa) rset = rdistinct rs rset @ rdistinct rsa (rset \<union> (set rs))"
Chengsong
parents: 478
diff changeset
   755
  apply(induct rs arbitrary: rsa rset)
478
Chengsong
parents: 476
diff changeset
   756
   apply simp
479
Chengsong
parents: 478
diff changeset
   757
  apply(case_tac "a \<in> rset")
Chengsong
parents: 478
diff changeset
   758
   apply simp
Chengsong
parents: 478
diff changeset
   759
  apply (simp add: insert_absorb)
Chengsong
parents: 478
diff changeset
   760
  by auto
478
Chengsong
parents: 476
diff changeset
   761
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   762
lemma middle_grewrites: 
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   763
"rs1 \<leadsto>g* rs2 \<Longrightarrow> rsa @ rs1 @ rsb \<leadsto>g* rsa @ rs2 @ rsb "
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   764
  by (simp add: grewritess_concat)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   765
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   766
lemma rdistinct_removes_all:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   767
  shows "set rs \<subseteq> rset \<Longrightarrow> rdistinct rs rset = []"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   768
  
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   769
  by (metis append.right_neutral rdistinct.simps(1) rdistinct_concat)
478
Chengsong
parents: 476
diff changeset
   770
479
Chengsong
parents: 478
diff changeset
   771
lemma ends_removal:
Chengsong
parents: 478
diff changeset
   772
  shows " rsb @ rdistinct rs (set rsb) @ RALTS rs # rsc \<leadsto>g* rsb @ rdistinct rs (set rsb) @ rsc"
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   773
  apply(subgoal_tac "rsb @ rdistinct rs (set rsb) @ RALTS rs # rsc \<leadsto>g* 
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   774
                     rsb @ rdistinct rs (set rsb) @ rs @ rsc")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   775
   apply(subgoal_tac "rsb @ rdistinct rs (set rsb) @ rs @ rsc \<leadsto>g* 
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   776
rsb @ rdistinct rs (set rsb) @ rdistinct rs (set (rsb @ rdistinct rs (set rsb))) @ rsc")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   777
  apply (metis (full_types) append_Nil2 append_eq_appendI greal_trans list_dlist_union rdistinct_removes_all set_append)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   778
   apply (metis append.assoc append_Nil gstar_rdistinct_general middle_grewrites)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   779
  using gr_in_rstar grewrite.intros(2) grewrites_append by presburger
478
Chengsong
parents: 476
diff changeset
   780
479
Chengsong
parents: 478
diff changeset
   781
lemma grewrites_rev_append:
Chengsong
parents: 478
diff changeset
   782
  shows "rs1 \<leadsto>g* rs2 \<Longrightarrow> rs1 @ [x] \<leadsto>g* rs2 @ [x]"
Chengsong
parents: 478
diff changeset
   783
  using grewritess_concat by auto
478
Chengsong
parents: 476
diff changeset
   784
479
Chengsong
parents: 478
diff changeset
   785
lemma grewrites_inclusion:
Chengsong
parents: 478
diff changeset
   786
  shows "set rs \<subseteq> set rs1 \<Longrightarrow> rs1 @ rs \<leadsto>g* rs1"
Chengsong
parents: 478
diff changeset
   787
  apply(induct rs arbitrary: rs1)
Chengsong
parents: 478
diff changeset
   788
  apply simp
Chengsong
parents: 478
diff changeset
   789
  by (meson gmany_steps_later grewrite_variant1 list.set_intros(1) set_subset_Cons subset_code(1))
478
Chengsong
parents: 476
diff changeset
   790
479
Chengsong
parents: 478
diff changeset
   791
lemma distinct_keeps_last:
Chengsong
parents: 478
diff changeset
   792
  shows "\<lbrakk>x \<notin> rset; x \<notin> set xs \<rbrakk> \<Longrightarrow> rdistinct (xs @ [x]) rset = rdistinct xs rset @ [x]"
Chengsong
parents: 478
diff changeset
   793
  by (simp add: concat_rdistinct_equality1)
478
Chengsong
parents: 476
diff changeset
   794
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   795
lemma grewrites_shape2_aux:
478
Chengsong
parents: 476
diff changeset
   796
  shows "       RALTS rs \<notin> set rsb \<Longrightarrow>
Chengsong
parents: 476
diff changeset
   797
       rsb @
Chengsong
parents: 476
diff changeset
   798
       rdistinct (rs @ rsa)
Chengsong
parents: 476
diff changeset
   799
        (set rsb) \<leadsto>g* rsb @
Chengsong
parents: 476
diff changeset
   800
                       rdistinct rs (set rsb) @
479
Chengsong
parents: 478
diff changeset
   801
                       rdistinct rsa (set rs \<union> set rsb \<union> {RALTS rs})"
Chengsong
parents: 478
diff changeset
   802
  apply(subgoal_tac " rdistinct (rs @ rsa) (set rsb) =  rdistinct rs (set rsb) @ rdistinct rsa (set rs \<union> set rsb)")
Chengsong
parents: 478
diff changeset
   803
   apply (simp only:)
Chengsong
parents: 478
diff changeset
   804
  prefer 2
Chengsong
parents: 478
diff changeset
   805
  apply (simp add: Un_commute concat_rdistinct_equality1)
Chengsong
parents: 478
diff changeset
   806
  apply(induct rsa arbitrary: rs rsb rule: rev_induct)
Chengsong
parents: 478
diff changeset
   807
   apply simp
Chengsong
parents: 478
diff changeset
   808
  apply(case_tac "x \<in> set rs")
Chengsong
parents: 478
diff changeset
   809
  apply (simp add: distinct_removes_middle3)
Chengsong
parents: 478
diff changeset
   810
  apply(case_tac "x = RALTS rs")
Chengsong
parents: 478
diff changeset
   811
   apply simp
Chengsong
parents: 478
diff changeset
   812
  apply(case_tac "x \<in> set rsb")
Chengsong
parents: 478
diff changeset
   813
   apply simp
Chengsong
parents: 478
diff changeset
   814
    apply (simp add: concat_rdistinct_equality1)
Chengsong
parents: 478
diff changeset
   815
  apply (simp add: concat_rdistinct_equality1)
Chengsong
parents: 478
diff changeset
   816
  apply simp
Chengsong
parents: 478
diff changeset
   817
  apply(drule_tac x = "rs " in meta_spec)
Chengsong
parents: 478
diff changeset
   818
  apply(drule_tac x = rsb in meta_spec)
Chengsong
parents: 478
diff changeset
   819
  apply simp
Chengsong
parents: 478
diff changeset
   820
  apply(subgoal_tac " rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \<union> set rsb) \<leadsto>g* rsb @ rdistinct rs (set rsb) @ rdistinct xs (insert (RALTS rs) (set rs \<union> set rsb))")
Chengsong
parents: 478
diff changeset
   821
  prefer 2
Chengsong
parents: 478
diff changeset
   822
   apply (simp add: concat_rdistinct_equality1)
Chengsong
parents: 478
diff changeset
   823
  apply(case_tac "x \<in> set xs")
Chengsong
parents: 478
diff changeset
   824
   apply simp
Chengsong
parents: 478
diff changeset
   825
   apply (simp add: distinct_removes_last2)
Chengsong
parents: 478
diff changeset
   826
  apply(case_tac "x \<in> set rsb")
Chengsong
parents: 478
diff changeset
   827
   apply (smt (verit, ccfv_threshold) Un_iff append.right_neutral concat_rdistinct_equality1 insert_is_Un rdistinct.simps(2))
Chengsong
parents: 478
diff changeset
   828
  apply(subgoal_tac "rsb @ rdistinct rs (set rsb) @ rdistinct (xs @ [x]) (set rs \<union> set rsb) = rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \<union> set rsb) @ [x]")
Chengsong
parents: 478
diff changeset
   829
  apply(simp only:)
Chengsong
parents: 478
diff changeset
   830
  apply(case_tac "x = RALTS rs")
Chengsong
parents: 478
diff changeset
   831
    apply(subgoal_tac "rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \<union> set rsb) @ [x] \<leadsto>g* rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \<union> set rsb) @ rs")
Chengsong
parents: 478
diff changeset
   832
  apply(subgoal_tac "rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \<union> set rsb) @ rs \<leadsto>g* rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \<union> set rsb) ")
Chengsong
parents: 478
diff changeset
   833
      apply (smt (verit, ccfv_SIG) Un_insert_left append.right_neutral concat_rdistinct_equality1 greal_trans insert_iff rdistinct.simps(2))
Chengsong
parents: 478
diff changeset
   834
  apply(subgoal_tac "set rs \<subseteq> set ( rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \<union> set rsb))")
Chengsong
parents: 478
diff changeset
   835
  apply (metis append.assoc grewrites_inclusion)
Chengsong
parents: 478
diff changeset
   836
     apply (metis Un_upper1 append.assoc dual_order.trans list_dlist_union set_append)
Chengsong
parents: 478
diff changeset
   837
  apply (metis append_Nil2 gr_in_rstar grewrite.intros(2) grewrite_append)
Chengsong
parents: 478
diff changeset
   838
   apply(subgoal_tac " rsb @ rdistinct rs (set rsb) @ rdistinct (xs @ [x]) (insert (RALTS rs) (set rs \<union> set rsb)) =  rsb @ rdistinct rs (set rsb) @ rdistinct (xs) (insert (RALTS rs) (set rs \<union> set rsb)) @ [x]")
Chengsong
parents: 478
diff changeset
   839
  apply(simp only:)
Chengsong
parents: 478
diff changeset
   840
  apply (metis append.assoc grewrites_rev_append)
Chengsong
parents: 478
diff changeset
   841
  apply (simp add: insert_absorb)
Chengsong
parents: 478
diff changeset
   842
   apply (simp add: distinct_keeps_last)+
Chengsong
parents: 478
diff changeset
   843
  done
478
Chengsong
parents: 476
diff changeset
   844
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   845
lemma grewrites_shape2:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   846
  shows "       RALTS rs \<notin> set rsb \<Longrightarrow>
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   847
       rsb @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   848
       rdistinct (rs @ rsa)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   849
        (set rsb) \<leadsto>g* rflts rsb @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   850
                       rdistinct rs (set rsb) @
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   851
                       rdistinct rsa (set rs \<union> set rsb \<union> {RALTS rs})"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   852
  apply (meson flts_gstar greal_trans grewrites.simps grewrites_shape2_aux grewritess_concat)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   853
  done
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   854
479
Chengsong
parents: 478
diff changeset
   855
lemma rdistinct_add_acc:
Chengsong
parents: 478
diff changeset
   856
  shows "rset2 \<subseteq> set rsb \<Longrightarrow> rsb @ rdistinct rs rset \<leadsto>g* rsb @ rdistinct rs (rset \<union> rset2)"
Chengsong
parents: 478
diff changeset
   857
  apply(induct rs arbitrary: rsb rset rset2)
Chengsong
parents: 478
diff changeset
   858
   apply simp
Chengsong
parents: 478
diff changeset
   859
  apply (case_tac "a \<in> rset")
Chengsong
parents: 478
diff changeset
   860
   apply simp
Chengsong
parents: 478
diff changeset
   861
  apply(case_tac "a \<in> rset2")
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   862
   apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   863
  apply (simp add: more_distinct1)
479
Chengsong
parents: 478
diff changeset
   864
  apply simp
Chengsong
parents: 478
diff changeset
   865
  apply(drule_tac x = "rsb @ [a]" in meta_spec)
Chengsong
parents: 478
diff changeset
   866
  by (metis Un_insert_left append.assoc append_Cons append_Nil set_append sup.coboundedI1)
Chengsong
parents: 478
diff changeset
   867
  
478
Chengsong
parents: 476
diff changeset
   868
479
Chengsong
parents: 478
diff changeset
   869
lemma frewrite_fun1:
Chengsong
parents: 478
diff changeset
   870
  shows "       RALTS rs \<in> set rsb \<Longrightarrow>
Chengsong
parents: 478
diff changeset
   871
       rsb @ rdistinct rsa (set rsb) \<leadsto>g* rflts rsb @ rdistinct rsa (set rsb \<union> set rs)"
Chengsong
parents: 478
diff changeset
   872
  apply(subgoal_tac "rsb @ rdistinct rsa (set rsb) \<leadsto>g* rflts rsb @ rdistinct rsa (set rsb)")
Chengsong
parents: 478
diff changeset
   873
   apply(subgoal_tac " set rs \<subseteq> set (rflts rsb)")
Chengsong
parents: 478
diff changeset
   874
  prefer 2
Chengsong
parents: 478
diff changeset
   875
  using spilled_alts_contained apply blast
Chengsong
parents: 478
diff changeset
   876
   apply(subgoal_tac "rflts rsb @ rdistinct rsa (set rsb) \<leadsto>g* rflts rsb @ rdistinct rsa (set rsb \<union> set rs)")
Chengsong
parents: 478
diff changeset
   877
  using greal_trans apply blast
Chengsong
parents: 478
diff changeset
   878
  using rdistinct_add_acc apply presburger
Chengsong
parents: 478
diff changeset
   879
  using flts_gstar grewritess_concat by auto
Chengsong
parents: 478
diff changeset
   880
  
478
Chengsong
parents: 476
diff changeset
   881
Chengsong
parents: 476
diff changeset
   882
Chengsong
parents: 476
diff changeset
   883
476
Chengsong
parents: 475
diff changeset
   884
Chengsong
parents: 475
diff changeset
   885
lemma frewrite_rd_grewrites:
Chengsong
parents: 475
diff changeset
   886
  shows "rs1 \<leadsto>f rs2 \<Longrightarrow> 
478
Chengsong
parents: 476
diff changeset
   887
\<exists>rs3. (rs @ (rdistinct rs1 (set rs)) \<leadsto>g* rs3) \<and> (rs @ (rdistinct rs2 (set rs)) \<leadsto>g* rs3) "
Chengsong
parents: 476
diff changeset
   888
  apply(induct rs1 rs2 arbitrary: rs rule: frewrite.induct)
Chengsong
parents: 476
diff changeset
   889
    apply(rule_tac x = "rsa @ (rdistinct rs ({RZERO} \<union> set rsa))" in exI)
476
Chengsong
parents: 475
diff changeset
   890
    apply(rule conjI)
478
Chengsong
parents: 476
diff changeset
   891
  apply(case_tac "RZERO \<in> set rsa")
Chengsong
parents: 476
diff changeset
   892
  apply simp+
Chengsong
parents: 476
diff changeset
   893
  using gstar0 apply fastforce
Chengsong
parents: 476
diff changeset
   894
     apply (simp add: gr_in_rstar grewrite.intros(1) grewrites_append)
Chengsong
parents: 476
diff changeset
   895
    apply (simp add: gstar0)
Chengsong
parents: 476
diff changeset
   896
    prefer 2
Chengsong
parents: 476
diff changeset
   897
    apply(case_tac "r \<in> set rs")
Chengsong
parents: 476
diff changeset
   898
  apply simp
Chengsong
parents: 476
diff changeset
   899
    apply(drule_tac x = "rs @ [r]" in meta_spec)
Chengsong
parents: 476
diff changeset
   900
    apply(erule exE)
Chengsong
parents: 476
diff changeset
   901
    apply(rule_tac x = "rs3" in exI)
Chengsong
parents: 476
diff changeset
   902
   apply simp
Chengsong
parents: 476
diff changeset
   903
  apply(case_tac "RALTS rs \<in> set rsb")
Chengsong
parents: 476
diff changeset
   904
   apply simp
479
Chengsong
parents: 478
diff changeset
   905
   apply(rule_tac x = "rflts rsb @ rdistinct rsa (set rsb \<union> set rs)" in exI)
478
Chengsong
parents: 476
diff changeset
   906
   apply(rule conjI)
479
Chengsong
parents: 478
diff changeset
   907
  using frewrite_fun1 apply force
Chengsong
parents: 478
diff changeset
   908
  apply (metis frewrite_fun1 rdistinct_concat sup_ge2)
478
Chengsong
parents: 476
diff changeset
   909
  apply(simp)
Chengsong
parents: 476
diff changeset
   910
  apply(rule_tac x = 
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   911
 "rflts rsb @
479
Chengsong
parents: 478
diff changeset
   912
                       rdistinct rs (set rsb) @
Chengsong
parents: 478
diff changeset
   913
                       rdistinct rsa (set rs \<union> set rsb \<union> {RALTS rs})" in exI)
478
Chengsong
parents: 476
diff changeset
   914
  apply(rule conjI)
Chengsong
parents: 476
diff changeset
   915
   prefer 2
Chengsong
parents: 476
diff changeset
   916
  using grewrites_shape2 apply force
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
   917
  using frewrite_rd_grewrites_aux by blast
478
Chengsong
parents: 476
diff changeset
   918
473
Chengsong
parents: 472
diff changeset
   919
476
Chengsong
parents: 475
diff changeset
   920
Chengsong
parents: 475
diff changeset
   921
lemma frewrites_rd_grewrites:
Chengsong
parents: 475
diff changeset
   922
  shows "rs1 \<leadsto>f* rs2 \<Longrightarrow>
478
Chengsong
parents: 476
diff changeset
   923
rsimp (RALTS rs1) = rsimp (RALTS rs2)"
476
Chengsong
parents: 475
diff changeset
   924
  apply(induct rs1 rs2 rule: frewrites.induct)
Chengsong
parents: 475
diff changeset
   925
   apply simp
479
Chengsong
parents: 478
diff changeset
   926
  using frewrite_simpeq by presburger
476
Chengsong
parents: 475
diff changeset
   927
Chengsong
parents: 475
diff changeset
   928
lemma frewrite_simpeq2:
Chengsong
parents: 475
diff changeset
   929
  shows "rs1 \<leadsto>f rs2 \<Longrightarrow> rsimp (RALTS (rdistinct rs1 {})) = rsimp (RALTS (rdistinct rs2 {}))"
478
Chengsong
parents: 476
diff changeset
   930
  apply(subgoal_tac "\<exists> rs3. (rdistinct rs1 {} \<leadsto>g* rs3) \<and> (rdistinct rs2 {} \<leadsto>g* rs3)")
Chengsong
parents: 476
diff changeset
   931
  using grewrites_equal_rsimp apply fastforce
479
Chengsong
parents: 478
diff changeset
   932
  by (metis append_self_conv2 frewrite_rd_grewrites list.set(1))
Chengsong
parents: 478
diff changeset
   933
Chengsong
parents: 478
diff changeset
   934
Chengsong
parents: 478
diff changeset
   935
478
Chengsong
parents: 476
diff changeset
   936
473
Chengsong
parents: 472
diff changeset
   937
(*a more refined notion of \<leadsto>* is needed,
Chengsong
parents: 472
diff changeset
   938
this lemma fails when rs1 contains some RALTS rs where elements
Chengsong
parents: 472
diff changeset
   939
of rs appear in later parts of rs1, which will be picked up by rs2
Chengsong
parents: 472
diff changeset
   940
and deduplicated*)
476
Chengsong
parents: 475
diff changeset
   941
lemma frewrites_simpeq:
473
Chengsong
parents: 472
diff changeset
   942
  shows "rs1 \<leadsto>f* rs2 \<Longrightarrow>
476
Chengsong
parents: 475
diff changeset
   943
 rsimp (RALTS (rdistinct rs1 {})) = rsimp (RALTS ( rdistinct rs2 {})) "
Chengsong
parents: 475
diff changeset
   944
  apply(induct rs1 rs2 rule: frewrites.induct)
Chengsong
parents: 475
diff changeset
   945
   apply simp
478
Chengsong
parents: 476
diff changeset
   946
  using frewrite_simpeq2 by presburger
476
Chengsong
parents: 475
diff changeset
   947
473
Chengsong
parents: 472
diff changeset
   948
Chengsong
parents: 472
diff changeset
   949
lemma frewrite_single_step:
Chengsong
parents: 472
diff changeset
   950
  shows " rs2 \<leadsto>f rs3 \<Longrightarrow> rsimp (RALTS rs2) = rsimp (RALTS rs3)"
Chengsong
parents: 472
diff changeset
   951
  apply(induct rs2 rs3 rule: frewrite.induct)
Chengsong
parents: 472
diff changeset
   952
    apply simp
Chengsong
parents: 472
diff changeset
   953
  using simp_flatten apply blast
Chengsong
parents: 472
diff changeset
   954
  by (metis (no_types, opaque_lifting) list.simps(9) rsimp.simps(2) simp_flatten2)
Chengsong
parents: 472
diff changeset
   955
Chengsong
parents: 472
diff changeset
   956
lemma frewrites_equivalent_simp:
Chengsong
parents: 472
diff changeset
   957
  shows "rs1 \<leadsto>f* rs2 \<Longrightarrow> rsimp (RALTS rs1) = rsimp (RALTS rs2)"
Chengsong
parents: 472
diff changeset
   958
  apply(induct rs1 rs2 rule: frewrites.induct)
Chengsong
parents: 472
diff changeset
   959
   apply simp
Chengsong
parents: 472
diff changeset
   960
  using frewrite_single_step by presburger
Chengsong
parents: 472
diff changeset
   961
476
Chengsong
parents: 475
diff changeset
   962
lemma grewrite_simpalts:
Chengsong
parents: 475
diff changeset
   963
  shows " rs2 \<leadsto>g rs3 \<Longrightarrow> rsimp (rsimp_ALTs rs2) = rsimp (rsimp_ALTs rs3)"
Chengsong
parents: 475
diff changeset
   964
  apply(induct rs2 rs3 rule : grewrite.induct)
Chengsong
parents: 475
diff changeset
   965
  using identity_wwo0 apply presburger
Chengsong
parents: 475
diff changeset
   966
  apply (metis frewrite.intros(1) frewrite_single_step identity_wwo0 rsimp_ALTs.simps(3) simp_flatten)
Chengsong
parents: 475
diff changeset
   967
  apply (smt (verit, ccfv_SIG) gmany_steps_later grewrites.intros(1) grewrites_cons grewrites_equal_rsimp identity_wwo0 rsimp_ALTs.simps(3))
Chengsong
parents: 475
diff changeset
   968
  apply simp
Chengsong
parents: 475
diff changeset
   969
  apply(subst rsimp_alts_equal)
Chengsong
parents: 475
diff changeset
   970
  apply(case_tac "rsa = [] \<and> rsb = [] \<and> rsc = []")
Chengsong
parents: 475
diff changeset
   971
   apply(subgoal_tac "rsa @ a # rsb @ rsc = [a]")
Chengsong
parents: 475
diff changeset
   972
  apply (simp only:)
Chengsong
parents: 475
diff changeset
   973
  apply (metis append_Nil frewrite.intros(1) frewrite_single_step identity_wwo0 rsimp_ALTs.simps(3) simp_removes_duplicate1(2))
Chengsong
parents: 475
diff changeset
   974
   apply simp
Chengsong
parents: 475
diff changeset
   975
  by (smt (verit, best) append.assoc append_Cons frewrite.intros(1) frewrite_single_step identity_wwo0 in_set_conv_decomp rsimp_ALTs.simps(3) simp_removes_duplicate3)
Chengsong
parents: 475
diff changeset
   976
Chengsong
parents: 475
diff changeset
   977
Chengsong
parents: 475
diff changeset
   978
lemma grewrites_simpalts:
Chengsong
parents: 475
diff changeset
   979
  shows " rs2 \<leadsto>g* rs3 \<Longrightarrow> rsimp (rsimp_ALTs rs2) = rsimp (rsimp_ALTs rs3)"
Chengsong
parents: 475
diff changeset
   980
  apply(induct rs2 rs3 rule: grewrites.induct)
Chengsong
parents: 475
diff changeset
   981
   apply simp
Chengsong
parents: 475
diff changeset
   982
  using grewrite_simpalts by presburger
473
Chengsong
parents: 472
diff changeset
   983
476
Chengsong
parents: 475
diff changeset
   984
473
Chengsong
parents: 472
diff changeset
   985
Chengsong
parents: 472
diff changeset
   986
Chengsong
parents: 472
diff changeset
   987
467
Chengsong
parents: 465
diff changeset
   988
lemma simp_der_flts:
471
Chengsong
parents: 469
diff changeset
   989
  shows "rsimp (RALTS (rdistinct (map (rder x) (rflts rs)) {})) = 
Chengsong
parents: 469
diff changeset
   990
         rsimp (RALTS (rdistinct (rflts (map (rder x) rs)) {}))"
473
Chengsong
parents: 472
diff changeset
   991
  apply(subgoal_tac "map (rder x) (rflts rs) \<leadsto>f* rflts (map (rder x) rs)")
478
Chengsong
parents: 476
diff changeset
   992
  using frewrites_simpeq apply presburger
Chengsong
parents: 476
diff changeset
   993
  using early_late_der_frewrites by auto
Chengsong
parents: 476
diff changeset
   994
Chengsong
parents: 476
diff changeset
   995
  
476
Chengsong
parents: 475
diff changeset
   996
Chengsong
parents: 475
diff changeset
   997
Chengsong
parents: 475
diff changeset
   998
Chengsong
parents: 475
diff changeset
   999
Chengsong
parents: 475
diff changeset
  1000
 
Chengsong
parents: 475
diff changeset
  1001
Chengsong
parents: 475
diff changeset
  1002
Chengsong
parents: 475
diff changeset
  1003
Chengsong
parents: 475
diff changeset
  1004
lemma simp_der_pierce_flts_prelim:
Chengsong
parents: 475
diff changeset
  1005
  shows "rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts rs)) {})) 
Chengsong
parents: 475
diff changeset
  1006
       = rsimp (rsimp_ALTs (rdistinct (rflts (map (rder x) rs)) {}))"
478
Chengsong
parents: 476
diff changeset
  1007
  by (metis append.right_neutral grewrite.intros(2) grewrite_simpalts rsimp_ALTs.simps(2) simp_der_flts)
467
Chengsong
parents: 465
diff changeset
  1008
Chengsong
parents: 465
diff changeset
  1009
465
Chengsong
parents: 456
diff changeset
  1010
lemma simp_der_pierce_flts:
471
Chengsong
parents: 469
diff changeset
  1011
  shows " rsimp (
Chengsong
parents: 469
diff changeset
  1012
rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {})
Chengsong
parents: 469
diff changeset
  1013
) =
Chengsong
parents: 469
diff changeset
  1014
          rsimp (
Chengsong
parents: 469
diff changeset
  1015
rsimp_ALTs (rdistinct (rflts (map (rder x) (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {})
Chengsong
parents: 469
diff changeset
  1016
)"
476
Chengsong
parents: 475
diff changeset
  1017
  using simp_der_pierce_flts_prelim by presburger
465
Chengsong
parents: 456
diff changeset
  1018
453
Chengsong
parents: 451
diff changeset
  1019
Chengsong
parents: 451
diff changeset
  1020
Chengsong
parents: 451
diff changeset
  1021
lemma simp_more_distinct:
465
Chengsong
parents: 456
diff changeset
  1022
  shows "rsimp  (rsimp_ALTs (rsa @ rs)) = rsimp (rsimp_ALTs (rsa @ (rdistinct rs (set rsa)))) "
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1023
  oops
465
Chengsong
parents: 456
diff changeset
  1024
453
Chengsong
parents: 451
diff changeset
  1025
Chengsong
parents: 451
diff changeset
  1026
lemma non_empty_list:
Chengsong
parents: 451
diff changeset
  1027
  shows "a \<in> set as \<Longrightarrow> as \<noteq> []"
Chengsong
parents: 451
diff changeset
  1028
  by (metis empty_iff empty_set)
Chengsong
parents: 451
diff changeset
  1029
456
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1030
lemma distinct_comp:
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1031
  shows "rdistinct (rs1@rs2) {} = (rdistinct rs1 {}) @ (rdistinct rs2 (set rs1))"
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1032
  apply(induct rs2 arbitrary: rs1)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1033
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1034
  apply(subgoal_tac "rs1 @ a # rs2 = (rs1 @ [a]) @ rs2")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1035
   apply(simp only:)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1036
   apply(case_tac "a \<in> set rs1")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1037
    apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1038
  oops
453
Chengsong
parents: 451
diff changeset
  1039
456
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1040
lemma instantiate1:
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1041
  shows "\<lbrakk>\<And>ab rset1.  rdistinct (ab # as) rset1 = rdistinct (ab # as @ [ab]) rset1\<rbrakk> \<Longrightarrow>  
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1042
rdistinct (aa # as) rset = rdistinct (aa # as @ [aa]) rset"
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1043
  apply(drule_tac x = "aa" in meta_spec)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1044
  apply(drule_tac x = "rset" in meta_spec)
453
Chengsong
parents: 451
diff changeset
  1045
  apply simp
456
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1046
  done
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1047
  
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1048
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1049
lemma not_head_elem:
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1050
  shows " \<lbrakk>aa \<in> set (a # as); aa \<notin> (set as)\<rbrakk> \<Longrightarrow> a = aa"
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1051
  
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1052
  by fastforce
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1053
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1054
(*
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1055
  apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1056
  apply (metis append_Cons)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1057
  apply(case_tac "ab \<in> rset1")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1058
  apply (metis (no_types, opaque_lifting) Un_insert_left append_Cons insert_iff rdistinct.simps(2) sup_bot_left)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1059
  apply(subgoal_tac "rdistinct (ab # (aa # as) @ [ab]) rset1 = 
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1060
               ab # (rdistinct ((aa # as) @ [ab]) (insert ab rset1))")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1061
   apply(simp only:)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1062
   apply(subgoal_tac "rdistinct (ab # aa # as) rset1 = ab # (rdistinct (aa # as) (insert ab rset1))")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1063
  apply(simp only:)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1064
    apply(subgoal_tac "rdistinct ((aa # as) @ [ab]) (insert ab rset1) = rdistinct (aa # as) (insert ab rset1)")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1065
  apply blast
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1066
*)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1067
  
453
Chengsong
parents: 451
diff changeset
  1068
Chengsong
parents: 451
diff changeset
  1069
lemma flts_identity1:
Chengsong
parents: 451
diff changeset
  1070
  shows  "rflts (rs @ [RONE]) = rflts rs @ [RONE] "
Chengsong
parents: 451
diff changeset
  1071
  apply(induct rs)
Chengsong
parents: 451
diff changeset
  1072
   apply simp+
Chengsong
parents: 451
diff changeset
  1073
  apply(case_tac a)
Chengsong
parents: 451
diff changeset
  1074
       apply simp
Chengsong
parents: 451
diff changeset
  1075
      apply simp+
Chengsong
parents: 451
diff changeset
  1076
  done
Chengsong
parents: 451
diff changeset
  1077
Chengsong
parents: 451
diff changeset
  1078
lemma flts_identity10:
Chengsong
parents: 451
diff changeset
  1079
  shows " rflts (rs @ [RCHAR c]) = rflts rs @ [RCHAR c]"
Chengsong
parents: 451
diff changeset
  1080
  apply(induct rs)
Chengsong
parents: 451
diff changeset
  1081
   apply simp+
Chengsong
parents: 451
diff changeset
  1082
  apply(case_tac a)
Chengsong
parents: 451
diff changeset
  1083
       apply simp+
Chengsong
parents: 451
diff changeset
  1084
  done
Chengsong
parents: 451
diff changeset
  1085
Chengsong
parents: 451
diff changeset
  1086
lemma flts_identity11:
Chengsong
parents: 451
diff changeset
  1087
  shows " rflts (rs @ [RSEQ r1 r2]) = rflts rs @ [RSEQ r1 r2]"
Chengsong
parents: 451
diff changeset
  1088
  apply(induct rs)
Chengsong
parents: 451
diff changeset
  1089
   apply simp+
Chengsong
parents: 451
diff changeset
  1090
  apply(case_tac a)
Chengsong
parents: 451
diff changeset
  1091
       apply simp+
Chengsong
parents: 451
diff changeset
  1092
  done
Chengsong
parents: 451
diff changeset
  1093
Chengsong
parents: 451
diff changeset
  1094
lemma flts_identity12:
Chengsong
parents: 451
diff changeset
  1095
  shows " rflts (rs @ [RSTAR r0]) = rflts rs @ [RSTAR r0]"
Chengsong
parents: 451
diff changeset
  1096
  apply(induct rs)
Chengsong
parents: 451
diff changeset
  1097
   apply simp+
Chengsong
parents: 451
diff changeset
  1098
  apply(case_tac a)
Chengsong
parents: 451
diff changeset
  1099
       apply simp+
Chengsong
parents: 451
diff changeset
  1100
  done
Chengsong
parents: 451
diff changeset
  1101
Chengsong
parents: 451
diff changeset
  1102
lemma flts_identity2:
Chengsong
parents: 451
diff changeset
  1103
  shows "a \<noteq> RZERO \<and> (\<forall>rs. a \<noteq> RALTS rs) \<Longrightarrow>  rflts (rs @ [a]) = rflts rs @ [a]"
Chengsong
parents: 451
diff changeset
  1104
  apply(case_tac a)
Chengsong
parents: 451
diff changeset
  1105
       apply simp
Chengsong
parents: 451
diff changeset
  1106
  using flts_identity1 apply auto[1]
Chengsong
parents: 451
diff changeset
  1107
  using flts_identity10 apply blast
Chengsong
parents: 451
diff changeset
  1108
  using flts_identity11 apply auto[1]
Chengsong
parents: 451
diff changeset
  1109
   apply blast
Chengsong
parents: 451
diff changeset
  1110
  using flts_identity12 by presburger
456
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1111
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1112
lemma flts_identity3:
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1113
  shows "a = RZERO  \<Longrightarrow> rflts (rs @ [a]) = rflts rs"
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1114
  apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1115
  apply(induct rs)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1116
   apply simp+
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1117
  apply(case_tac aa)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1118
       apply simp+
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1119
  done
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1120
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1121
lemma distinct_removes_last3:
465
Chengsong
parents: 456
diff changeset
  1122
  shows "\<lbrakk>a \<in> set as\<rbrakk>
456
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1123
    \<Longrightarrow> rdistinct as {} = rdistinct (as @ [a]) {}"
465
Chengsong
parents: 456
diff changeset
  1124
  by (simp add: distinct_removes_last2)
456
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1125
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1126
lemma set_inclusion_with_flts1:
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1127
  shows " \<lbrakk>RONE \<in> set rs\<rbrakk> \<Longrightarrow> RONE  \<in> set (rflts rs)"
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1128
  apply(induct rs)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1129
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1130
  apply(case_tac " RONE \<in> set rs")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1131
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1132
  apply (metis Un_upper2 insert_absorb insert_subset list.set_intros(2) rflts.simps(2) rflts.simps(3) rflts_def_idiot set_append)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1133
  apply(case_tac "RONE = a")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1134
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1135
  apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1136
  done
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1137
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1138
lemma set_inclusion_with_flts10:
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1139
  shows " \<lbrakk>RCHAR x \<in> set rs\<rbrakk> \<Longrightarrow> RCHAR x  \<in> set (rflts rs)"
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1140
  apply(induct rs)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1141
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1142
  apply(case_tac " RCHAR x \<in> set rs")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1143
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1144
  apply (metis Un_upper2 insert_absorb insert_subset rflts.simps(2) rflts.simps(3) rflts_def_idiot set_append set_subset_Cons)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1145
  apply(case_tac "RCHAR x = a")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1146
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1147
  apply fastforce
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1148
  apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1149
  done
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1150
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1151
lemma set_inclusion_with_flts11:
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1152
  shows " \<lbrakk>RSEQ r1 r2 \<in> set rs\<rbrakk> \<Longrightarrow> RSEQ r1 r2  \<in> set (rflts rs)"
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1153
  apply(induct rs)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1154
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1155
  apply(case_tac " RSEQ r1 r2 \<in> set rs")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1156
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1157
  apply (metis Un_upper2 insert_absorb insert_subset rflts.simps(2) rflts.simps(3) rflts_def_idiot set_append set_subset_Cons)
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1158
  apply(case_tac "RSEQ r1 r2 = a")
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1159
   apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1160
  apply fastforce
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1161
  apply simp
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1162
  done
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1163
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1164
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1165
lemma set_inclusion_with_flts:
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1166
  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))"
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1167
  by (simp add: set_inclusion_with_flts1)
453
Chengsong
parents: 451
diff changeset
  1168
  
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1169
lemma can_spill_lst:"\<And>x5. \<lbrakk>a \<in> set as; rsimp a \<in> set (map rsimp as); rsimp a = RALTS x5\<rbrakk>
456
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1170
          \<Longrightarrow> rsimp_ALTs (rdistinct (rflts (map rsimp as @ [rsimp a])) {}) = 
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1171
rsimp_ALTs (rdistinct (rflts (map rsimp as @ x5)) {})"
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1172
  
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1173
  using flts_append rflts_spills_last rsimp_inner_idem4 by presburger
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1174
  
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1175
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1176
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1177
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1178
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1179
lemma common_rewrites_equal:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1180
  shows "(rs1 \<leadsto>g* rs3) \<and> (rs2 \<leadsto>g* rs3) \<Longrightarrow> rsimp (rsimp_ALTs rs1 ) = rsimp (rsimp_ALTs rs2)"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1181
  using grewrites_simpalts by force
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1182
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1183
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1184
lemma basic_regex_property1:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1185
  shows "rnullable r \<Longrightarrow> rsimp r \<noteq> RZERO"
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1186
  apply(induct r rule: rsimp.induct)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1187
  apply(auto)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1188
  apply (metis idiot idiot2 rrexp.distinct(5))
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1189
  by (metis der_simp_nullability rnullable.simps(1) rnullable.simps(4) rsimp.simps(2))
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1190
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1191
lemma basic_rsimp_SEQ_property1:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1192
  shows "rsimp_SEQ RONE r = r"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1193
  by (simp add: idiot)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1194
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1195
lemma basic_rsimp_SEQ_property3:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1196
  shows "rsimp_SEQ r RZERO = RZERO"  
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1197
  using rsimp_SEQ.elims by blast
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1198
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1199
thm rsimp_SEQ.elims
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1200
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1201
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1202
lemma basic_rsimp_SEQ_property2:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1203
  shows "\<lbrakk>r1 \<noteq> RZERO ; r1 \<noteq> RONE; r2 \<noteq> RZERO\<rbrakk> \<Longrightarrow>rsimp_SEQ r1 r2 = RSEQ r1 r2"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1204
  apply(case_tac r1)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1205
       apply simp+
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1206
     apply (simp add: idiot2)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1207
  using idiot2 apply blast
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1208
  using idiot2 apply auto[1]
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1209
  using idiot2 by blast
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1210
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1211
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1212
(*
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1213
lemma rderssimp_same_rewrites_rder_induct1:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1214
  shows "\<lbrakk> ([rder x (rsimp r1)] \<leadsto>g* rs1) \<and> ([rder x r1] \<leadsto>g* rs1) ;
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1215
 ([rder x (rsimp r2)] \<leadsto>g* rs2) \<and> ([rder x r2] \<leadsto>g* rs2) \<rbrakk> \<Longrightarrow> 
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1216
\<exists>rs3. ([rder x (rsimp (RSEQ r1 r2))] \<leadsto>g* rs3) \<and> ([rder x (RSEQ r1 r2)] \<leadsto>g* rs3) "
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1217
  sorry
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1218
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1219
lemma rderssimp_same_rewrites_rder_induct2:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1220
  shows "\<lbrakk> ([rder x (rsimp r1)] \<leadsto>g* rs1) \<and> ([rder x r1] \<leadsto>g* rs1) \<rbrakk> \<Longrightarrow> 
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1221
\<exists>rs3. ([rder x (rsimp (RSTAR r1))] \<leadsto>g* rs3) \<and> ([rder x (RSTAR r1)] \<leadsto>g* rs3) "
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1222
  sorry
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1223
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1224
lemma rderssimp_same_rewrites_rder_induct3:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1225
  shows "\<lbrakk> ([rder x (rsimp r1)] \<leadsto>g* rs1) \<and> ([rder x r1] \<leadsto>g* rs1) ;
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1226
 ([rder x (rsimp r2)] \<leadsto>g* rs2) \<and> ([rder x r2] \<leadsto>g* rs2) \<rbrakk> \<Longrightarrow> 
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1227
\<exists>rs3. ([rder x (rsimp (RALT r1 r2))] \<leadsto>g* rs3) \<and> ([rder x (RALT r1 r2)] \<leadsto>g* rs3) "
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1228
  sorry
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1229
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1230
lemma rderssimp_same_rewrites_rder_induct4:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1231
  shows "\<lbrakk>\<forall>r \<in> set rs. \<exists> rsa. ([rder x (rsimp r)] \<leadsto>g* rsa ) \<and> ([rder x r] \<leadsto>g* rsa) \<rbrakk> \<Longrightarrow> 
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1232
\<exists>rsb. ([rder x (rsimp (RALTS rs))]  \<leadsto>g* rsb) \<and> ([rder x (RALTS rs)] \<leadsto>g* rsb) "
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1233
  sorry
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1234
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1235
lemma rderssimp_same_rewrites_rder_base1:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1236
  shows "([rder x (rsimp RONE)] \<leadsto>g* [] ) \<and> ([rder x RONE] \<leadsto>g* [])"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1237
  by (simp add: gr_in_rstar grewrite.intros(1))
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1238
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1239
lemma rderssimp_same_rewrites_rder_base2:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1240
  shows " ([rder x (rsimp RZERO)] \<leadsto>g* [] ) \<and> ([rder x RZERO] \<leadsto>g* [])"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1241
  using rderssimp_same_rewrites_rder_base1 by auto
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1242
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1243
lemma rderssimp_same_rewrites_rder_base3:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1244
  shows " ([rder x (rsimp (RCHAR c))] \<leadsto>g* [] ) \<and> ([rder x (RCHAR c)] \<leadsto>g* [])"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1245
  sorry
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1246
*)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1247
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1248
lemma inside_simp_seq_nullable:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1249
  shows 
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1250
"\<And>r1 r2.
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1251
       \<lbrakk>rsimp (rder x (rsimp r1)) = rsimp (rder x r1); rsimp (rder x (rsimp r2)) = rsimp (rder x r2);
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1252
        rnullable r1\<rbrakk>
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1253
       \<Longrightarrow> rsimp (rder x (rsimp_SEQ (rsimp r1) (rsimp r2))) =
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1254
           rsimp_ALTs (rdistinct (rflts [rsimp_SEQ (rsimp (rder x r1)) (rsimp r2), rsimp (rder x r2)]) {})"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1255
  apply(case_tac "rsimp r1 = RONE")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1256
   apply(simp)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1257
  apply(subst basic_rsimp_SEQ_property1)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1258
   apply (simp add: idem_after_simp1)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1259
  apply(case_tac "rsimp r1 = RZERO")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1260
  
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1261
  using basic_regex_property1 apply blast
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1262
  apply(case_tac "rsimp r2 = RZERO")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1263
  
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1264
  apply (simp add: basic_rsimp_SEQ_property3)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1265
  apply(subst basic_rsimp_SEQ_property2)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1266
     apply simp+
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1267
  apply(subgoal_tac "rnullable (rsimp r1)")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1268
   apply simp
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1269
  using rsimp_idem apply presburger
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1270
  using der_simp_nullability by presburger
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1271
  
456
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1272
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1273
inductive 
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1274
  hrewrite:: "rrexp \<Rightarrow> rrexp \<Rightarrow> bool" ("_ \<leadsto> _" [99, 99] 99)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1275
where
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1276
  "RSEQ  RZERO r2 \<leadsto> RZERO"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1277
| "RSEQ  r1 RZERO \<leadsto> RZERO"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1278
| "RSEQ  RONE r \<leadsto>  r"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1279
| "r1 \<leadsto> r2 \<Longrightarrow> RSEQ  r1 r3 \<leadsto> RSEQ r2 r3"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1280
| "r3 \<leadsto> r4 \<Longrightarrow> RSEQ r1 r3 \<leadsto> RSEQ r1 r4"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1281
| "r \<leadsto> r' \<Longrightarrow> (RALTS (rs1 @ [r] @ rs2)) \<leadsto> (RALTS  (rs1 @ [r'] @ rs2))"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1282
(*context rule for eliminating 0, alts--corresponds to the recursive call flts r::rs = r::(flts rs)*)
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1283
| "RALTS  (rsa @ [RZERO] @ rsb) \<leadsto> RALTS  (rsa @ rsb)"
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1284
| "RALTS  (rsa @ [RALTS rs1] @ rsb) \<leadsto> RALTS (rsa @ rs1 @ rsb)"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1285
| "RALTS  [] \<leadsto> RZERO"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1286
| "RALTS  [r] \<leadsto> r"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1287
| "a1 = a2 \<Longrightarrow> RALTS (rsa@[a1]@rsb@[a2]@rsc) \<leadsto> RALTS (rsa @ [a1] @ rsb @ rsc)"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1288
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1289
inductive 
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1290
  hrewrites:: "rrexp \<Rightarrow> rrexp \<Rightarrow> bool" ("_ \<leadsto>* _" [100, 100] 100)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1291
where 
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1292
  rs1[intro, simp]:"r \<leadsto>* r"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1293
| rs2[intro]: "\<lbrakk>r1 \<leadsto>* r2; r2 \<leadsto> r3\<rbrakk> \<Longrightarrow> r1 \<leadsto>* r3"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1294
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1295
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1296
lemma hr_in_rstar : "r1 \<leadsto> r2 \<Longrightarrow> r1 \<leadsto>* r2"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1297
  using hrewrites.intros(1) hrewrites.intros(2) by blast
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1298
 
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1299
lemma hreal_trans[trans]: 
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1300
  assumes a1: "r1 \<leadsto>* r2"  and a2: "r2 \<leadsto>* r3"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1301
  shows "r1 \<leadsto>* r3"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1302
  using a2 a1
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1303
  apply(induct r2 r3 arbitrary: r1 rule: hrewrites.induct) 
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1304
  apply(auto)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1305
  done
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1306
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1307
lemma  hmany_steps_later: "\<lbrakk>r1 \<leadsto> r2; r2 \<leadsto>* r3 \<rbrakk> \<Longrightarrow> r1 \<leadsto>* r3"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1308
  by (meson hr_in_rstar hreal_trans)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1309
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1310
lemma hrewrites_seq_context:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1311
  shows "r1 \<leadsto>* r2 \<Longrightarrow> RSEQ r1 r3 \<leadsto>* RSEQ r2 r3"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1312
  apply(induct r1 r2 rule: hrewrites.induct)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1313
   apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1314
  using hrewrite.intros(4) by blast
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1315
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1316
lemma hrewrites_seq_context2:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1317
  shows "r1 \<leadsto>* r2 \<Longrightarrow> RSEQ r0 r1 \<leadsto>* RSEQ r0 r2"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1318
  apply(induct r1 r2 rule: hrewrites.induct)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1319
   apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1320
  using hrewrite.intros(5) by blast
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1321
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1322
lemma hrewrites_seq_context0:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1323
  shows "r1 \<leadsto>* RZERO \<Longrightarrow> RSEQ r1 r3 \<leadsto>* RZERO"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1324
  apply(subgoal_tac "RSEQ r1 r3 \<leadsto>* RSEQ RZERO r3")
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1325
  using hrewrite.intros(1) apply blast
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1326
  by (simp add: hrewrites_seq_context)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1327
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1328
lemma hrewrites_seq_contexts:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1329
  shows "\<lbrakk>r1 \<leadsto>* r2; r3 \<leadsto>* r4\<rbrakk> \<Longrightarrow> RSEQ r1 r3 \<leadsto>* RSEQ r2 r4"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1330
  by (meson hreal_trans hrewrites_seq_context hrewrites_seq_context2)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1331
488
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1332
lemma hrewrite_simpeq:
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1333
  shows "r1 \<leadsto> r2 \<Longrightarrow> rsimp r1 = rsimp r2"
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1334
  apply(induct rule: hrewrite.induct)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1335
            apply simp+
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1336
  apply (simp add: basic_rsimp_SEQ_property3)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1337
  apply (simp add: basic_rsimp_SEQ_property1)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1338
  using rsimp.simps(1) apply presburger
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1339
        apply simp+
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1340
  using flts_middle0 apply force
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1341
  using simp_flatten3 apply presburger
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1342
  apply simp+
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1343
  apply (simp add: idem_after_simp1)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1344
  using grewrite.intros(4) grewrite_equal_rsimp by presburger
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1345
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1346
lemma hrewrites_simpeq:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1347
  shows "r1 \<leadsto>* r2 \<Longrightarrow> rsimp r1 = rsimp r2"
488
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1348
  apply(induct rule: hrewrites.induct)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1349
   apply simp
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1350
  apply(subgoal_tac "rsimp r2 = rsimp r3")
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1351
   apply auto[1]
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1352
  using hrewrite_simpeq by presburger
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1353
  
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1354
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1355
lemma grewrite_ralts:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1356
  shows "rs \<leadsto>g rs' \<Longrightarrow> RALTS rs \<leadsto>* RALTS rs'"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1357
  by (smt (verit) grewrite_cases_middle hr_in_rstar hrewrite.intros(11) hrewrite.intros(7) hrewrite.intros(8))
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1358
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1359
lemma grewrites_ralts:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1360
  shows "rs \<leadsto>g* rs' \<Longrightarrow> RALTS rs \<leadsto>* RALTS rs'"
488
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1361
  apply(induct rule: grewrites.induct)
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1362
  apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1363
  using grewrite_ralts hreal_trans by blast
488
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1364
  
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1365
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1366
lemma distinct_grewrites_subgoal1:
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1367
  shows "  
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1368
       \<lbrakk>rs1 \<leadsto>g* [a]; RALTS rs1 \<leadsto>* a; [a] \<leadsto>g rs3\<rbrakk> \<Longrightarrow> RALTS rs1 \<leadsto>* rsimp_ALTs rs3"
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1369
  apply(subgoal_tac "RALTS rs1 \<leadsto>* RALTS rs3")
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1370
  apply (metis hrewrite.intros(10) hrewrite.intros(9) rs2 rsimp_ALTs.cases rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3))
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1371
  apply(subgoal_tac "rs1 \<leadsto>g* rs3")
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1372
  using grewrites_ralts apply blast
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1373
  using grewrites.intros(2) by presburger
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1374
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1375
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1376
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1377
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1378
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1379
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1380
lemma grewrites_ralts_rsimpalts:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1381
  shows "rs \<leadsto>g* rs' \<Longrightarrow> RALTS rs \<leadsto>* rsimp_ALTs rs' "
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1382
  apply(induct rs rs' rule: grewrites.induct)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1383
   apply(case_tac rs)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1384
  using hrewrite.intros(9) apply force
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1385
   apply(case_tac list)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1386
  apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1387
  using hr_in_rstar hrewrite.intros(10) rsimp_ALTs.simps(2) apply presburger
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1388
   apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1389
  apply(case_tac rs2)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1390
   apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1391
   apply (metis grewrite.intros(3) grewrite_singleton rsimp_ALTs.simps(1))
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1392
  apply(case_tac list)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1393
   apply(simp)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1394
  using distinct_grewrites_subgoal1 apply blast
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1395
  apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1396
  apply(case_tac rs3)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1397
   apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1398
  using grewrites_ralts hrewrite.intros(9) apply blast
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1399
  by (metis (no_types, opaque_lifting) grewrite_ralts hr_in_rstar hreal_trans hrewrite.intros(10) neq_Nil_conv rsimp_ALTs.simps(2) rsimp_ALTs.simps(3))
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1400
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1401
lemma hrewrites_alts:
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1402
  shows " r \<leadsto>* r' \<Longrightarrow> (RALTS (rs1 @ [r] @ rs2)) \<leadsto>* (RALTS  (rs1 @ [r'] @ rs2))"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1403
  apply(induct r r' rule: hrewrites.induct)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1404
  apply simp
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1405
  using hrewrite.intros(6) by blast
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1406
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1407
inductive 
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1408
  srewritescf:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" (" _ scf\<leadsto>* _" [100, 100] 100)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1409
where
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1410
  ss1: "[] scf\<leadsto>* []"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1411
| ss2: "\<lbrakk>r \<leadsto>* r'; rs scf\<leadsto>* rs'\<rbrakk> \<Longrightarrow> (r#rs) scf\<leadsto>* (r'#rs')"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1412
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1413
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1414
lemma hrewrites_alts_cons:
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1415
  shows " RALTS rs \<leadsto>* RALTS rs' \<Longrightarrow> RALTS (r # rs) \<leadsto>* RALTS (r # rs')"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1416
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1417
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1418
  oops
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1419
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1420
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1421
lemma srewritescf_alt: "rs1 scf\<leadsto>* rs2 \<Longrightarrow> (RALTS (rs@rs1)) \<leadsto>* (RALTS (rs@rs2))"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1422
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1423
  apply(induct rs1 rs2 arbitrary: rs rule: srewritescf.induct)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1424
   apply(rule rs1)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1425
  apply(drule_tac x = "rsa@[r']" in meta_spec)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1426
  apply simp
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1427
  apply(rule hreal_trans)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1428
   prefer 2
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1429
   apply(assumption)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1430
  apply(drule hrewrites_alts)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1431
  by auto
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1432
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1433
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1434
corollary srewritescf_alt1: 
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1435
  assumes "rs1 scf\<leadsto>* rs2"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1436
  shows "RALTS rs1 \<leadsto>* RALTS rs2"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1437
  using assms
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1438
  by (metis append_Nil srewritescf_alt)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1439
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1440
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1441
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1442
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1443
lemma trivialrsimp_srewrites: 
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1444
  "\<lbrakk>\<And>x. x \<in> set rs \<Longrightarrow> x \<leadsto>* f x \<rbrakk> \<Longrightarrow> rs scf\<leadsto>* (map f rs)"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1445
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1446
  apply(induction rs)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1447
   apply simp
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1448
   apply(rule ss1)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1449
  by (metis insert_iff list.simps(15) list.simps(9) srewritescf.simps)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1450
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1451
lemma hrewrites_list: 
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1452
  shows
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1453
" (\<And>xa. xa \<in> set x \<Longrightarrow> xa \<leadsto>* rsimp xa) \<Longrightarrow> RALTS x \<leadsto>* RALTS (map rsimp x)"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1454
  apply(induct x)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1455
   apply(simp)+
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1456
  by (simp add: srewritescf_alt1 ss2 trivialrsimp_srewrites)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1457
(*  apply(subgoal_tac "RALTS x \<leadsto>* RALTS (map rsimp x)")*)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1458
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1459
  
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1460
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1461
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1462
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1463
lemma simp_hrewrites:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1464
  shows "r1 \<leadsto>* rsimp r1"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1465
  apply(induct r1)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1466
       apply simp+
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1467
    apply(case_tac "rsimp r11 = RONE")
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1468
     apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1469
     apply(subst basic_rsimp_SEQ_property1)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1470
  apply(subgoal_tac "RSEQ r11 r12 \<leadsto>* RSEQ RONE r12")
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1471
  using hreal_trans hrewrite.intros(3) apply blast
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1472
  using hrewrites_seq_context apply presburger
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1473
    apply(case_tac "rsimp r11 = RZERO")
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1474
     apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1475
  using hrewrite.intros(1) hrewrites_seq_context apply blast
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1476
    apply(case_tac "rsimp r12 = RZERO")
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1477
     apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1478
  apply(subst basic_rsimp_SEQ_property3)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1479
  apply (meson hrewrite.intros(2) hrewrites.simps hrewrites_seq_context2)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1480
    apply(subst basic_rsimp_SEQ_property2)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1481
       apply simp+
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1482
  using hrewrites_seq_contexts apply presburger
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1483
   apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1484
   apply(subgoal_tac "RALTS x \<leadsto>* RALTS (map rsimp x)")
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1485
  apply(subgoal_tac "RALTS (map rsimp x) \<leadsto>* rsimp_ALTs (rdistinct (rflts (map rsimp x)) {}) ")
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1486
  using hreal_trans apply blast
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1487
    apply (meson flts_gstar greal_trans grewrites_ralts_rsimpalts gstar_rdistinct)
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1488
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1489
   apply (simp add: grewrites_ralts hrewrites_list)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1490
  by simp
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1491
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1492
lemma interleave_aux1:
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1493
  shows " RALT (RSEQ RZERO r1) r \<leadsto>*  r"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1494
  apply(subgoal_tac "RSEQ RZERO r1 \<leadsto>* RZERO")
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1495
  apply(subgoal_tac "RALT (RSEQ RZERO r1) r \<leadsto>* RALT RZERO r")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1496
  apply (meson grewrite.intros(1) grewrite_ralts hreal_trans hrewrite.intros(10) hrewrites.simps)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1497
  using rs1 srewritescf_alt1 ss1 ss2 apply presburger
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1498
  by (simp add: hr_in_rstar hrewrite.intros(1))
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1499
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1500
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1501
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1502
lemma rnullable_hrewrite:
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1503
  shows "r1 \<leadsto> r2 \<Longrightarrow> rnullable r1 = rnullable r2"
488
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1504
  apply(induct rule: hrewrite.induct)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1505
            apply simp+
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1506
     apply blast
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1507
    apply simp+
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1508
  done
465
Chengsong
parents: 456
diff changeset
  1509
453
Chengsong
parents: 451
diff changeset
  1510
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1511
lemma interleave1:
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1512
  shows "r \<leadsto> r' \<Longrightarrow> rder c r \<leadsto>* rder c r'"
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1513
  apply(induct r r' rule: hrewrite.induct)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1514
            apply (simp add: hr_in_rstar hrewrite.intros(1))
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1515
  apply (metis (no_types, lifting) basic_rsimp_SEQ_property3 list.simps(8) list.simps(9) rder.simps(1) rder.simps(5) rdistinct.simps(1) rflts.simps(1) rflts.simps(2) rsimp.simps(1) rsimp.simps(2) rsimp.simps(3) rsimp_ALTs.simps(1) simp_hrewrites)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1516
          apply simp
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1517
          apply(subst interleave_aux1)
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1518
          apply simp
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1519
         apply(case_tac "rnullable r1")
488
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1520
          apply simp
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1521
  
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1522
          apply (simp add: hrewrites_seq_context rnullable_hrewrite srewritescf_alt1 ss1 ss2)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1523
  
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1524
         apply (simp add: hrewrites_seq_context rnullable_hrewrite)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1525
        apply(case_tac "rnullable r1")
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1526
  apply simp
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1527
  
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1528
  using hr_in_rstar hrewrites_seq_context2 srewritescf_alt1 ss1 ss2 apply presburger
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1529
  apply simp
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1530
  using hr_in_rstar hrewrites_seq_context2 apply blast
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1531
       apply simp
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1532
  
488
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1533
  using hrewrites_alts apply auto[1]
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1534
  apply simp
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1535
  using grewrite.intros(1) grewrite_append grewrite_ralts apply auto[1]
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1536
  apply simp
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1537
  apply (simp add: grewrite.intros(2) grewrite_append grewrite_ralts)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1538
  apply (simp add: hr_in_rstar hrewrite.intros(9))
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1539
   apply (simp add: hr_in_rstar hrewrite.intros(10))
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1540
  apply simp
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1541
  using hrewrite.intros(11) by auto
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1542
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1543
lemma interleave_star1:
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1544
  shows "r \<leadsto>* r' \<Longrightarrow> rder c r \<leadsto>* rder c r'"
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1545
  apply(induct rule : hrewrites.induct)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1546
   apply simp
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1547
  by (meson hreal_trans interleave1)
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1548
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1549
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1550
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1551
lemma inside_simp_removal:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1552
  shows " rsimp (rder x (rsimp r)) = rsimp (rder x r)"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1553
  apply(induct r)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1554
       apply simp+
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1555
    apply(case_tac "rnullable r1")
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1556
     apply simp
456
26a5e640cdd7 realPhdThesis
Chengsong
parents: 453
diff changeset
  1557
  
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1558
  using inside_simp_seq_nullable apply blast
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1559
    apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1560
  apply (smt (verit, del_insts) basic_rsimp_SEQ_property2 basic_rsimp_SEQ_property3 der_simp_nullability rder.simps(1) rder.simps(5) rnullable.simps(2) rsimp.simps(1) rsimp_SEQ.simps(1) rsimp_idem)
488
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1561
   apply(subgoal_tac "rder x (RALTS xa) \<leadsto>* rder x (rsimp (RALTS xa))")
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1562
  using hrewrites_simpeq apply presburger
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1563
  using interleave_star1 simp_hrewrites apply presburger
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  1564
  by simp
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1565
  
451
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1566
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1567
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1568
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1569
lemma rders_simp_same_simpders:
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1570
  shows "s \<noteq> [] \<Longrightarrow> rders_simp r s = rsimp (rders r s)"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1571
  apply(induct s rule: rev_induct)
453
Chengsong
parents: 451
diff changeset
  1572
   apply simp
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1573
  apply(case_tac "xs = []")
453
Chengsong
parents: 451
diff changeset
  1574
   apply simp
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1575
  apply(simp add: rders_append rders_simp_append)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1576
  using inside_simp_removal by blast
453
Chengsong
parents: 451
diff changeset
  1577
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1578
453
Chengsong
parents: 451
diff changeset
  1579
Chengsong
parents: 451
diff changeset
  1580
451
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1581
lemma distinct_der:
480
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1582
  shows "rsimp (rsimp_ALTs (map (rder x) (rdistinct rs {}))) = 
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1583
         rsimp (rsimp_ALTs (rdistinct (map (rder x) rs) {}))"
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1584
  by (metis grewrites_simpalts gstar_rdistinct inside_simp_removal rder_rsimp_ALTs_commute)
574749f5190b a few more
Chengsong
parents: 479
diff changeset
  1585
451
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1586
453
Chengsong
parents: 451
diff changeset
  1587
  
Chengsong
parents: 451
diff changeset
  1588
Chengsong
parents: 451
diff changeset
  1589
Chengsong
parents: 451
diff changeset
  1590
lemma rders_simp_lambda:
Chengsong
parents: 451
diff changeset
  1591
  shows " rsimp \<circ> rder x \<circ> (\<lambda>r. rders_simp r xs) = (\<lambda>r. rders_simp r (xs @ [x]))"
Chengsong
parents: 451
diff changeset
  1592
  using rders_simp_append by auto
451
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1593
453
Chengsong
parents: 451
diff changeset
  1594
lemma rders_simp_nonempty_simped:
Chengsong
parents: 451
diff changeset
  1595
  shows "xs \<noteq> [] \<Longrightarrow> rsimp \<circ> (\<lambda>r. rders_simp r xs) = (\<lambda>r. rders_simp r xs)"
Chengsong
parents: 451
diff changeset
  1596
  using rders_simp_same_simpders rsimp_idem by auto
Chengsong
parents: 451
diff changeset
  1597
Chengsong
parents: 451
diff changeset
  1598
lemma repeated_altssimp:
Chengsong
parents: 451
diff changeset
  1599
  shows "\<forall>r \<in> set rs. rsimp r = r \<Longrightarrow> rsimp (rsimp_ALTs (rdistinct (rflts rs) {})) =
Chengsong
parents: 451
diff changeset
  1600
           rsimp_ALTs (rdistinct (rflts rs) {})"
Chengsong
parents: 451
diff changeset
  1601
  by (metis map_idI rsimp.simps(2) rsimp_idem)
451
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1602
465
Chengsong
parents: 456
diff changeset
  1603
Chengsong
parents: 456
diff changeset
  1604
lemma add0_isomorphic:
Chengsong
parents: 456
diff changeset
  1605
  shows "rsimp_ALTs (rdistinct (rflts [rsimp r, RZERO]) {}) = rsimp r"
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1606
  by (metis append.left_neutral append_Cons flts_removes0 idem_after_simp1)
465
Chengsong
parents: 456
diff changeset
  1607
Chengsong
parents: 456
diff changeset
  1608
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1609
465
Chengsong
parents: 456
diff changeset
  1610
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1611
lemma alts_closed_form: shows 
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1612
"rsimp (rders_simp (RALTS rs) s) = 
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1613
rsimp (RALTS (map (\<lambda>r. rders_simp r s) rs))"
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1614
  apply(induct s rule: rev_induct)
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1615
   apply simp
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1616
  apply simp
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1617
  apply(subst rders_simp_append)
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1618
  apply(subgoal_tac " rsimp (rders_simp (rders_simp (RALTS rs) xs) [x]) = 
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1619
 rsimp(rders_simp (rsimp_ALTs (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {})) [x])")
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1620
   prefer 2
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1621
  apply (metis inside_simp_removal rders_simp_one_char)
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1622
  apply(simp only: )
451
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1623
  apply(subst rders_simp_one_char)
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1624
  apply(subst rsimp_idem)
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1625
  apply(subgoal_tac "rsimp (rder x (rsimp_ALTs (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {}))) =
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1626
                     rsimp ((rsimp_ALTs (map (rder x) (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {})))) ")
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1627
  prefer 2
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1628
  using rder_rsimp_ALTs_commute apply presburger
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1629
  apply(simp only:)
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1630
  apply(subgoal_tac "rsimp (rsimp_ALTs (map (rder x) (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {})))
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1631
= rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {}))")
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1632
   prefer 2
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1633
  
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1634
  using distinct_der apply presburger
7a016eeb118d finiteness
Chengsong
parents: 445
diff changeset
  1635
  apply(simp only:)
453
Chengsong
parents: 451
diff changeset
  1636
  apply(subgoal_tac " rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {})) =
Chengsong
parents: 451
diff changeset
  1637
                      rsimp (rsimp_ALTs (rdistinct ( (rflts (map (rder x) (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)))) {}))")
Chengsong
parents: 451
diff changeset
  1638
   apply(simp only:)
Chengsong
parents: 451
diff changeset
  1639
  apply(subgoal_tac " rsimp (rsimp_ALTs (rdistinct (rflts (map (rder x) (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {})) = 
Chengsong
parents: 451
diff changeset
  1640
                      rsimp (rsimp_ALTs (rdistinct (rflts ( (map (rsimp \<circ> (rder x) \<circ> (\<lambda>r. rders_simp r xs)) rs))) {}))")
Chengsong
parents: 451
diff changeset
  1641
    apply(simp only:)
Chengsong
parents: 451
diff changeset
  1642
  apply(subst rders_simp_lambda)
Chengsong
parents: 451
diff changeset
  1643
    apply(subst rders_simp_nonempty_simped)
Chengsong
parents: 451
diff changeset
  1644
     apply simp
Chengsong
parents: 451
diff changeset
  1645
    apply(subgoal_tac "\<forall>r \<in> set  (map (\<lambda>r. rders_simp r (xs @ [x])) rs). rsimp r = r")
Chengsong
parents: 451
diff changeset
  1646
  prefer 2
Chengsong
parents: 451
diff changeset
  1647
     apply (simp add: rders_simp_same_simpders rsimp_idem)
Chengsong
parents: 451
diff changeset
  1648
    apply(subst repeated_altssimp)
Chengsong
parents: 451
diff changeset
  1649
     apply simp
Chengsong
parents: 451
diff changeset
  1650
  apply fastforce
465
Chengsong
parents: 456
diff changeset
  1651
   apply (metis inside_simp_removal list.map_comp rder.simps(4) rsimp.simps(2) rsimp_idem)
476
Chengsong
parents: 475
diff changeset
  1652
  using simp_der_pierce_flts by blast
443
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1653
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1654
lemma alts_closed_form_variant: shows 
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1655
"s \<noteq> [] \<Longrightarrow> rders_simp (RALTS rs) s = 
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1656
rsimp (RALTS (map (\<lambda>r. rders_simp r s) rs))"
478
Chengsong
parents: 476
diff changeset
  1657
  by (metis alts_closed_form comp_apply rders_simp_nonempty_simped)
443
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1658
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1659
thm vsuf.simps
443
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1660
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1661
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1662
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1663
lemma rsimp_seq_equal1:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1664
  shows "rsimp_SEQ (rsimp r1) (rsimp r2) = rsimp_ALTs (rdistinct (rflts [rsimp_SEQ (rsimp r1) (rsimp r2)]) {})"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1665
  by (metis idem_after_simp1 rsimp.simps(1))
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1666
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1667
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1668
483
Chengsong
parents: 482
diff changeset
  1669
Chengsong
parents: 482
diff changeset
  1670
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1671
fun sflat_aux :: "rrexp  \<Rightarrow> rrexp list " where
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1672
  "sflat_aux (RALTS (r # rs)) = sflat_aux r @ rs"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1673
| "sflat_aux (RALTS []) = []"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1674
| "sflat_aux r = [r]"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1675
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1676
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1677
fun sflat :: "rrexp \<Rightarrow> rrexp" where
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1678
  "sflat (RALTS (r # [])) = r"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1679
| "sflat (RALTS (r # rs)) = RALTS (sflat_aux r @ rs)"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1680
| "sflat r = r"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1681
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1682
inductive created_by_seq:: "rrexp \<Rightarrow> bool" where
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1683
  "created_by_seq (RSEQ r1 r2) "
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1684
| "created_by_seq r1 \<Longrightarrow> created_by_seq (RALT r1 r2)"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1685
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1686
(*Why \<and> rnullable case would be provable.....?*)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1687
lemma seq_der_shape:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1688
  shows "\<forall>r1 r2. \<exists>r3 r4. (rder c (RSEQ r1 r2) = RSEQ r3 r4 \<or> rder c (RSEQ r1 r2) = RALT r3 r4)"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1689
  by (meson rder.simps(5))
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1690
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1691
lemma alt_der_shape:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1692
  shows "\<forall>rs. \<exists> rs'. (rder c (RALTS rs)) = RALTS (map (rder c) rs)"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1693
  by force
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1694
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1695
lemma seq_ders_shape1:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1696
  shows "\<forall>r1 r2. \<exists>r3 r4. (rders (RSEQ r1 r2) s) = RSEQ r3 r4 \<or> (rders (RSEQ r1 r2) s) = RALT r3 r4"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1697
  apply(induct s rule: rev_induct)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1698
   apply auto[1]
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1699
  apply(rule allI)+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1700
  apply(subst rders_append)+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1701
  apply(subgoal_tac " \<exists>r3 r4. rders (RSEQ r1 r2) xs = RSEQ r3 r4 \<or> rders (RSEQ r1 r2) xs = RALT r3 r4 ")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1702
   apply(erule exE)+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1703
   apply(erule disjE)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1704
    apply simp+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1705
  done
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1706
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1707
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1708
lemma seq_ders_shape2:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1709
  shows "rders (RSEQ r1 r2) s = RALTS (ra # rb # rs) \<Longrightarrow> rs = []"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1710
  using seq_ders_shape1
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1711
  by (metis list.inject rrexp.distinct(25) rrexp.inject(3))
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1712
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1713
lemma created_by_seq_der:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1714
  shows "created_by_seq r \<Longrightarrow> created_by_seq (rder c r)"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1715
  apply(induct r)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1716
  apply simp+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1717
  
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1718
  using created_by_seq.cases apply blast
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1719
  
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1720
  apply (meson created_by_seq.cases rrexp.distinct(19) rrexp.distinct(21))
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1721
  apply (metis created_by_seq.simps rder.simps(5))
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1722
   apply (smt (verit, ccfv_threshold) created_by_seq.simps list.set_intros(1) list.simps(8) list.simps(9) rder.simps(4) rrexp.distinct(25) rrexp.inject(3))
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1723
  using created_by_seq.intros(1) by force
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1724
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1725
lemma createdbyseq_left_creatable:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1726
  shows "created_by_seq (RALT r1 r2) \<Longrightarrow> created_by_seq r1"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1727
  using created_by_seq.cases by blast
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1728
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1729
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1730
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1731
lemma recursively_derseq:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1732
  shows " created_by_seq (rders (RSEQ r1 r2) s)"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1733
  apply(induct s rule: rev_induct)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1734
   apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1735
  using created_by_seq.intros(1) apply force
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1736
  apply(subgoal_tac "created_by_seq (rders (RSEQ r1 r2) (xs @ [x]))")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1737
  apply blast
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1738
  apply(subst rders_append)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1739
  apply(subgoal_tac "\<exists>r3 r4. rders (RSEQ r1 r2) xs = RSEQ r3 r4 \<or> 
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1740
                    rders (RSEQ r1 r2) xs = RALT r3 r4")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1741
   prefer 2
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1742
  using seq_ders_shape1 apply presburger
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1743
  apply(erule exE)+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1744
  apply(erule disjE)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1745
   apply(subgoal_tac "created_by_seq (rders (RSEQ r3 r4) [x])")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1746
    apply presburger
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1747
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1748
  using created_by_seq.intros(1) created_by_seq.intros(2) apply presburger
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1749
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1750
  apply(subgoal_tac "created_by_seq r3")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1751
  prefer 2
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1752
  using createdbyseq_left_creatable apply blast
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1753
  using created_by_seq.intros(2) created_by_seq_der by blast
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1754
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1755
  
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1756
lemma recursively_derseq1:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1757
  shows "r = rders (RSEQ r1 r2) s \<Longrightarrow> created_by_seq r"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1758
  using recursively_derseq by blast
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1759
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1760
lemma recursively_derseq2:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1761
  shows "r = rders (RSEQ r1 r2) s \<Longrightarrow> \<exists>r1 r2. r = RSEQ r1 r2 \<or> (r = RALT r1 r2 \<and> created_by_seq r1) "
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1762
  by (meson created_by_seq.cases recursively_derseq)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1763
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1764
lemma recursively_derseq3:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1765
  shows "created_by_seq r \<Longrightarrow> \<exists>r1 r2. r = RSEQ r1 r2 \<or> (r = RALT r1 r2 \<and> created_by_seq r1) "
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1766
  by (meson created_by_seq.cases recursively_derseq)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1767
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1768
lemma createdbyseq_prop1:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1769
  shows "created_by_seq (RALTS xs) \<Longrightarrow> \<exists> ra rb. xs = [ra, rb]"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1770
  by (metis created_by_seq.cases rrexp.inject(3) rrexp.simps(30))
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1771
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1772
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1773
lemma sfau_head:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1774
  shows " created_by_seq r \<Longrightarrow> \<exists>ra rb rs. sflat_aux r = RSEQ ra rb # rs"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1775
  apply(induction r rule: created_by_seq.induct)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1776
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1777
  by fastforce
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1778
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1779
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1780
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1781
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1782
lemma sfaux_eq10:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1783
  shows "created_by_seq r3 \<Longrightarrow> 
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1784
rs \<noteq> [] \<Longrightarrow> sflat (RALTS ((map (rder c) (sflat_aux r3)) @ (map (rder c) rs))) = 
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1785
RALTS (sflat_aux (rder c r3) @ (map (rder c ) rs) )"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1786
  apply(induction r3 arbitrary: rs rule: created_by_seq.induct)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1787
   apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1788
   apply(case_tac "rnullable r1")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1789
    apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1790
  
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1791
  apply (metis append_Cons list.exhaust map_is_Nil_conv self_append_conv2 sflat.simps(2) sflat_aux.simps(1) sflat_aux.simps(6))
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1792
   apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1793
  
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1794
  apply (metis Nil_is_map_conv append_Cons append_Nil list.exhaust sflat.simps(2) sflat_aux.simps(6))
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1795
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1796
  by force
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1797
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1798
lemma sfaux_keeps_tail:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1799
  shows "created_by_seq r3 \<Longrightarrow>  
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1800
         sflat_aux (RALTS (map (rder c) (sflat_aux r3) @ xs )) = 
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1801
         sflat_aux (RALTS (map (rder c) (sflat_aux r3))) @ xs  "
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1802
  using sfau_head by fastforce
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1803
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1804
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1805
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1806
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1807
lemma sfaux_eq00:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1808
  shows "created_by_seq r3 \<Longrightarrow> 
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1809
 sflat_aux (RALTS ((map (rder c) (sflat_aux r3)) @ (map (rder c) rs))) = 
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1810
 (sflat_aux (rder c r3) @ (map (rder c ) rs) )"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1811
  apply(induct rs rule: rev_induct)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1812
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1813
  apply (smt (verit, del_insts) append_Cons created_by_seq.simps list.distinct(1) list.simps(9) map_append rder.simps(4) rrexp.inject(3) self_append_conv self_append_conv2 sfau_head sfaux_eq10 sflat_aux.simps(1) sflat_aux.simps(6))
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1814
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1815
  apply(subst sfaux_keeps_tail)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1816
   apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1817
  apply(subst (asm) sfaux_keeps_tail)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1818
  apply simp+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1819
  done
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1820
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1821
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1822
lemma sfaux_eq1:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1823
  shows "created_by_seq r3 \<Longrightarrow> sflat (RALTS ((map (rder c) (sflat_aux r3)) @ [rder c r4])) = RALTS (sflat_aux (rder c r3) @ [rder c r4] )"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1824
  by (metis (no_types, lifting) list.distinct(1) list.simps(9) map_append map_concat_cons self_append_conv sfaux_eq10)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1825
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1826
lemma sflat_seq_induct:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1827
  shows "sflat (rder c (sflat (rders (RSEQ r1 r2) s) )) = sflat (rders (RSEQ r1 r2) (s @ [c]))"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1828
  apply(subst rders_append)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1829
  apply(case_tac "rders (RSEQ r1 r2) s ")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1830
  prefer 6
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1831
       apply simp+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1832
  apply(subgoal_tac "\<exists>r3 r4. x5 = [r3, r4]")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1833
  apply(erule exE)+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1834
   apply(subgoal_tac "sflat (rder c (sflat (RALT r3 r4))) = sflat (RALTS (map (rder c) [r3, r4]))")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1835
    apply meson
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1836
   apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1837
  
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1838
  apply (metis createdbyseq_left_creatable recursively_derseq sfaux_eq1)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1839
  by (metis created_by_seq.simps recursively_derseq rrexp.distinct(25) rrexp.inject(3))
484
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1840
15f02ec4d9fe changeto sflat def
Chengsong
parents: 483
diff changeset
  1841
483
Chengsong
parents: 482
diff changeset
  1842
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1843
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1844
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1845
lemma vsuf_prop1:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1846
  shows "vsuf (xs @ [x]) r = (if (rnullable (rders r xs)) 
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1847
                                then [x] # (map (\<lambda>s. s @ [x]) (vsuf xs r) )
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1848
                                else (map (\<lambda>s. s @ [x]) (vsuf xs r)) ) 
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1849
             "
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1850
  apply(induct xs arbitrary: r)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1851
   apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1852
  apply(case_tac "rnullable r")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1853
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1854
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1855
  done
483
Chengsong
parents: 482
diff changeset
  1856
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1857
fun  breakHead :: "rrexp list \<Rightarrow> rrexp list" where
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1858
  "breakHead [] = [] "
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1859
| "breakHead (RALT r1 r2 # rs) = r1 # r2 # rs"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1860
| "breakHead (r # rs) = r # rs"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1861
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1862
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1863
lemma sfau_idem_der:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1864
  shows "created_by_seq r \<Longrightarrow> sflat_aux (rder c r) = breakHead (map (rder c) (sflat_aux r))"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1865
  apply(induct rule: created_by_seq.induct)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1866
   apply simp+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1867
  using sfau_head by fastforce
483
Chengsong
parents: 482
diff changeset
  1868
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1869
lemma vsuf_compose1:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1870
  shows " \<not> rnullable (rders r1 xs)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1871
       \<Longrightarrow> map (rder x \<circ> rders r2) (vsuf xs r1) = map (rders r2) (vsuf (xs @ [x]) r1)"
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1872
  apply(subst vsuf_prop1)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1873
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1874
  by (simp add: rders_append)
483
Chengsong
parents: 482
diff changeset
  1875
  
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1876
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1877
483
Chengsong
parents: 482
diff changeset
  1878
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1879
lemma seq_sfau0:
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1880
  shows  "sflat_aux (rders (RSEQ r1 r2) s) = (RSEQ (rders r1 s) r2) #
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1881
                                       (map (rders r2) (vsuf s r1)) "
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1882
  apply(induct s rule: rev_induct)
483
Chengsong
parents: 482
diff changeset
  1883
   apply simp
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1884
  apply(subst rders_append)+
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1885
  apply(subgoal_tac "created_by_seq (rders (RSEQ r1 r2) xs)")
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1886
  prefer 2
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1887
  using recursively_derseq1 apply blast
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1888
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1889
  apply(subst sfau_idem_der)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1890
  
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1891
   apply blast
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1892
  apply(case_tac "rnullable (rders r1 xs)")
483
Chengsong
parents: 482
diff changeset
  1893
   apply simp
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1894
   apply(subst vsuf_prop1)
483
Chengsong
parents: 482
diff changeset
  1895
  apply simp
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1896
  apply (simp add: rders_append)
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1897
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1898
  using vsuf_compose1 by blast
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1899
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1900
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1901
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1902
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1903
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1904
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1905
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1906
483
Chengsong
parents: 482
diff changeset
  1907
Chengsong
parents: 482
diff changeset
  1908
thm sflat.elims
Chengsong
parents: 482
diff changeset
  1909
Chengsong
parents: 482
diff changeset
  1910
Chengsong
parents: 482
diff changeset
  1911
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1912
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1913
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1914
lemma sflat_rsimpeq:
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1915
  shows "created_by_seq r1 \<Longrightarrow> sflat_aux r1 =  rs \<Longrightarrow> rsimp r1 = rsimp (RALTS rs)"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1916
  apply(induct r1 arbitrary: rs rule:  created_by_seq.induct)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1917
   apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1918
  using rsimp_seq_equal1 apply force
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1919
  by (metis head_one_more_simp rsimp.simps(2) sflat_aux.simps(1) simp_flatten)
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1920
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1921
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1922
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1923
lemma seq_closed_form_general:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1924
  shows "rsimp (rders (RSEQ r1 r2) s) = 
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1925
rsimp ( (RALTS ( (RSEQ (rders r1 s) r2 # (map (rders r2) (vsuf s r1))))))"
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1926
  apply(case_tac "s \<noteq> []")
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1927
  apply(subgoal_tac "created_by_seq (rders (RSEQ r1 r2) s)")
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1928
  apply(subgoal_tac "sflat_aux (rders (RSEQ r1 r2) s) = RSEQ (rders r1 s) r2 # (map (rders r2) (vsuf s r1))")
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1929
  using sflat_rsimpeq apply blast
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1930
    apply (simp add: seq_sfau0)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1931
  using recursively_derseq1 apply blast
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1932
  apply simp
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1933
  by (metis idem_after_simp1 rsimp.simps(1))
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1934
  
485
72edbac05c59 central lemma for seqclosedforms
Chengsong
parents: 484
diff changeset
  1935
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  1936
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1937
lemma seq_closed_form_aux1:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1938
  shows "rsimp ( (RALTS ( (RSEQ (rders r1 s) r2 # (map (rders r2) (vsuf s r1)))))) =
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1939
rsimp ( (RALTS ( (RSEQ (rders_simp r1 s) r2 # (map (rders r2) (vsuf s r1))))))"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1940
  by (smt (verit, best) list.simps(9) rders.simps(1) rders_simp.simps(1) rders_simp_same_simpders rsimp.simps(1) rsimp.simps(2) rsimp_idem)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1941
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1942
lemma add_simp_to_rest:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1943
  shows "rsimp (RALTS (r # rs)) = rsimp (RALTS (r # map rsimp rs))"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1944
  by (metis append_Nil2 grewrite.intros(2) grewrite_simpalts head_one_more_simp list.simps(9) rsimp_ALTs.simps(2) spawn_simp_rsimpalts)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1945
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1946
lemma rsimp_compose_der:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1947
  shows "map rsimp (map (rders r) ss) = map (\<lambda>s. rsimp (rders r s)) ss"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1948
  apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1949
  done
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1950
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1951
lemma rsimp_compose_der2:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1952
  shows "\<forall>s \<in> set ss. s \<noteq> [] \<Longrightarrow> map rsimp (map (rders r) ss) = map (\<lambda>s.  (rders_simp r s)) ss"  
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1953
  by (simp add: rders_simp_same_simpders)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1954
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1955
lemma vsuf_nonempty:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1956
  shows "\<forall>s \<in> set ( vsuf s1 r). s \<noteq> []"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1957
  apply(induct s1 arbitrary: r)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1958
   apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1959
  apply simp
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1960
  done
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1961
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1962
lemma rsimp_compose_der3:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1963
  shows " map rsimp (map (rders r) (vsuf s1 r')) = map (\<lambda>s. rsimp  (rders_simp r s)) (vsuf s1 r')"  
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1964
  by (simp add: rders_simp_same_simpders rsimp_idem vsuf_nonempty)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1965
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1966
thm rders_simp_same_simpders
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1967
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1968
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1969
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1970
lemma seq_closed_form_aux2:
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1971
  shows "s \<noteq> [] \<Longrightarrow> rsimp ( (RALTS ( (RSEQ (rders_simp r1 s) r2 # (map (rders r2) (vsuf s r1)))))) = 
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1972
         rsimp ( (RALTS ( (RSEQ (rders_simp r1 s) r2 # (map (rders_simp r2) (vsuf s r1))))))"
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1973
  
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1974
  by (metis add_simp_to_rest rsimp_compose_der2 vsuf_nonempty)
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1975
  
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1976
443
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1977
lemma seq_closed_form: shows 
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1978
"rsimp (rders_simp (RSEQ r1 r2) s) = 
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1979
rsimp ( RALTS ( (RSEQ (rders_simp r1 s) r2) # 
445
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  1980
                (map (rders_simp r2) (vsuf s r1)) 
443
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1981
              )  
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1982
      )"
482
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1983
  apply(case_tac s )
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1984
   apply simp  
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1985
  apply (metis idem_after_simp1 rsimp.simps(1))
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1986
  apply(subgoal_tac " s \<noteq> []")
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1987
  using rders_simp_same_simpders rsimp_idem seq_closed_form_aux1 seq_closed_form_aux2 seq_closed_form_general apply presburger
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1988
  by fastforce
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1989
  
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1990
4bf2367e6e53 closedformseq
Chengsong
parents: 480
diff changeset
  1991
443
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1992
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  1993
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1994
lemma seq_closed_form_variant: shows
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1995
"s \<noteq> [] \<Longrightarrow> (rders_simp (RSEQ r1 r2) s) = 
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  1996
rsimp (RALTS ((RSEQ (rders_simp r1 s) r2) # (map (rders_simp r2) (vsuf s r1))))"
445
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  1997
  apply(induct s rule: rev_induct)
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  1998
   apply simp
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  1999
  apply(subst rders_simp_append)
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  2000
  apply(subst rders_simp_one_char)
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  2001
  apply(subst rsimp_idem[symmetric])
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  2002
  apply(subst rders_simp_one_char[symmetric])
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  2003
  apply(subst rders_simp_append[symmetric])
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  2004
  apply(insert seq_closed_form)
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  2005
  apply(subgoal_tac "rsimp (rders_simp (RSEQ r1 r2) (xs @ [x]))
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  2006
 = rsimp (RALTS (RSEQ (rders_simp r1 (xs @ [x])) r2 # map (rders_simp r2) (vsuf (xs @ [x]) r1)))")
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  2007
   apply force
e072cfc2f2ee closedforms
Chengsong
parents: 444
diff changeset
  2008
  by presburger
443
c6351a96bf80 writeupforclosedforms
Chengsong
parents:
diff changeset
  2009
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2010
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2011
lemma star_closed_form_seq1:
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2012
  shows "sflat_aux (rders (RSTAR r0) (c # s)) = 
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2013
   RSEQ (rders (rder c r0) s) (RSTAR r0) #
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2014
                                       map (rders (RSTAR r0)) (vsuf s (rder c r0))"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2015
  apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2016
  using seq_sfau0 by blast
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2017
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2018
fun hflat_aux :: "rrexp \<Rightarrow> rrexp list" where
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2019
  "hflat_aux (RALT r1 r2) = hflat_aux r1 @ hflat_aux r2"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2020
| "hflat_aux r = [r]"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2021
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2022
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2023
fun hflat :: "rrexp \<Rightarrow> rrexp" where
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2024
  "hflat (RALT r1 r2) = RALTS ((hflat_aux r1) @ (hflat_aux r2))"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2025
| "hflat r = r"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2026
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2027
inductive created_by_star :: "rrexp \<Rightarrow> bool" where
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2028
  "created_by_star (RSEQ ra (RSTAR rb))"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2029
| "\<lbrakk>created_by_star r1; created_by_star r2\<rbrakk> \<Longrightarrow> created_by_star (RALT r1 r2)"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2030
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2031
fun hElem :: "rrexp  \<Rightarrow> rrexp list" where
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2032
  "hElem (RALT r1 r2) = (hElem r1 ) @ (hElem r2)"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2033
| "hElem r = [r]"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2034
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2035
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2036
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2037
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2038
lemma cbs_ders_cbs:
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2039
  shows "created_by_star r \<Longrightarrow> created_by_star (rder c r)"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2040
  apply(induct r rule: created_by_star.induct)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2041
   apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2042
  using created_by_star.intros(1) created_by_star.intros(2) apply auto[1]
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2043
  by (metis (mono_tags, lifting) created_by_star.simps list.simps(8) list.simps(9) rder.simps(4))
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2044
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2045
lemma star_ders_cbs:
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2046
  shows "created_by_star (rders (RSEQ r1 (RSTAR r2)) s)"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2047
  apply(induct s rule: rev_induct)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2048
   apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2049
   apply (simp add: created_by_star.intros(1))
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2050
  apply(subst rders_append)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2051
  apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2052
  using cbs_ders_cbs by auto
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2053
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2054
lemma created_by_star_cases:
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2055
  shows "created_by_star r \<Longrightarrow> \<exists>ra rb. (r = RALT ra rb \<and> created_by_star ra \<and> created_by_star rb) \<or> r = RSEQ ra rb "
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2056
  by (meson created_by_star.cases)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2057
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2058
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2059
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2060
lemma hfau_pushin: 
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2061
  shows "created_by_star r \<Longrightarrow> hflat_aux (rder c r) = concat (map hElem (map (rder c) (hflat_aux r)))"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2062
  apply(induct r rule: created_by_star.induct)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2063
   apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2064
  apply(subgoal_tac "created_by_star (rder c r1)")
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2065
  prefer 2
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2066
  apply(subgoal_tac "created_by_star (rder c r2)")
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2067
  using cbs_ders_cbs apply blast
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2068
  using cbs_ders_cbs apply auto[1]
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2069
  apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2070
  done
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2071
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2072
lemma stupdate_induct1:
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2073
  shows " concat (map (hElem \<circ> (rder x \<circ> (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)))) Ss) =
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2074
          map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_update x r0 Ss)"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2075
  apply(induct Ss)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2076
   apply simp+
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2077
  by (simp add: rders_append)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2078
  
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2079
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2080
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2081
lemma stupdates_join_general:
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2082
  shows  "concat (map hElem (map (rder x) (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates xs r0 Ss)))) =
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2083
           map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates (xs @ [x]) r0 Ss)"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2084
  apply(induct xs arbitrary: Ss)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2085
   apply (simp)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2086
  prefer 2
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2087
   apply auto[1]
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2088
  using stupdate_induct1 by blast
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2089
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2090
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2091
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2092
 
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2093
lemma stupdates_join:
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2094
  shows "concat (map hElem (map (rder x) (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates xs r0 [[c]])))) =
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2095
           map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates (xs @ [x]) r0 [[c]])"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2096
  using stupdates_join_general by auto
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2097
  
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2098
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2099
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2100
lemma star_hfau_induct:
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2101
  shows "hflat_aux (rders (RSEQ (rder c r0) (RSTAR r0)) s) =   
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2102
      map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates s r0 [[c]])"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2103
  apply(induct s rule: rev_induct)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2104
   apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2105
  apply(subst rders_append)+
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2106
  apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2107
  apply(subst stupdates_append)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2108
  apply(subgoal_tac "created_by_star (rders (RSEQ (rder c r0) (RSTAR r0)) xs)")
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2109
  prefer 2
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2110
  apply (simp add: star_ders_cbs)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2111
  apply(subst hfau_pushin)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2112
   apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2113
  apply(subgoal_tac "concat (map hElem (map (rder x) (hflat_aux (rders (RSEQ (rder c r0) (RSTAR r0)) xs)))) =
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2114
                     concat (map hElem (map (rder x) ( map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates xs r0 [[c]])))) ")
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2115
   apply(simp only:)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2116
  prefer 2
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2117
   apply presburger
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2118
  apply(subst stupdates_append[symmetric])
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2119
  using stupdates_join_general by blast
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2120
487
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2121
lemma starders_hfau_also1:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2122
  shows "hflat_aux (rders (RSTAR r) (c # xs)) = map (\<lambda>s1. RSEQ (rders r s1) (RSTAR r)) (star_updates xs r [[c]])"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2123
  using star_hfau_induct by force
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2124
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2125
lemma hflat_aux_grewrites:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2126
  shows "a # rs \<leadsto>g* hflat_aux a @ rs"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2127
  apply(induct a arbitrary: rs)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2128
       apply simp+
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2129
   apply(case_tac x)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2130
    apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2131
  apply(case_tac list)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2132
  
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2133
  apply (metis append.right_neutral append_Cons append_eq_append_conv2 grewrites.simps hflat_aux.simps(7) same_append_eq)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2134
   apply(case_tac lista)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2135
  apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2136
  apply (metis (no_types, lifting) append_Cons append_eq_append_conv2 gmany_steps_later greal_trans grewrite.intros(2) grewrites_append self_append_conv)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2137
  apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2138
  by simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2139
  
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2140
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2141
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2142
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2143
lemma cbs_hfau_rsimpeq1:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2144
  shows "rsimp (RALT a b) = rsimp (RALTS ((hflat_aux a) @ (hflat_aux b)))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2145
  apply(subgoal_tac "[a, b] \<leadsto>g* hflat_aux a @ hflat_aux b")
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2146
  using grewrites_equal_rsimp apply presburger
488
370dae790b30 more sorrys fileld
Chengsong
parents: 487
diff changeset
  2147
  by (metis append.right_neutral greal_trans grewrites_cons hflat_aux_grewrites)
487
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2148
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2149
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2150
lemma hfau_rsimpeq2:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2151
  shows "created_by_star r \<Longrightarrow> rsimp r = rsimp ( (RALTS (hflat_aux r)))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2152
  apply(induct r)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2153
       apply simp+
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2154
  
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2155
    apply (metis rsimp_seq_equal1)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2156
  prefer 2
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2157
   apply simp
487
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2158
  apply(case_tac x)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2159
   apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2160
  apply(case_tac "list")
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2161
   apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2162
  
487
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2163
  apply (metis idem_after_simp1)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2164
  apply(case_tac "lista")
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2165
  prefer 2
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2166
   apply (metis hflat_aux.simps(8) idem_after_simp1 list.simps(8) list.simps(9) rsimp.simps(2))
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2167
  apply(subgoal_tac "rsimp (RALT a aa) = rsimp (RALTS (hflat_aux (RALT a aa)))")
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2168
  apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2169
  apply(subgoal_tac "rsimp (RALT a aa) = rsimp (RALTS (hflat_aux a @ hflat_aux aa))")
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2170
  using hflat_aux.simps(1) apply presburger
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2171
  apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2172
  using cbs_hfau_rsimpeq1 by fastforce
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2173
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2174
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2175
lemma hfau_rsimpeq1:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2176
  shows "created_by_star r \<Longrightarrow> hflat_aux r = rs \<Longrightarrow> rsimp r = rsimp (RALTS rs)"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2177
  apply(induct r arbitrary: rs rule: created_by_star.induct )
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2178
  apply (metis created_by_seq.intros(1) hflat_aux.simps(5) sflat_aux.simps(6) sflat_rsimpeq)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2179
  apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2180
  oops
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2181
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2182
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2183
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2184
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2185
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2186
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2187
lemma star_closed_form1:
487
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2188
  shows "rsimp (rders (RSTAR r0) (c#s)) = 
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2189
rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2190
  using hfau_rsimpeq2 rder.simps(6) rders.simps(2) star_ders_cbs starders_hfau_also1 by presburger
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2191
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2192
lemma star_closed_form2:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2193
  shows  "rsimp (rders_simp (RSTAR r0) (c#s)) = 
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2194
rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2195
  by (metis list.distinct(1) rders_simp_same_simpders rsimp_idem star_closed_form1)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2196
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2197
lemma star_closed_form3:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2198
  shows  "rsimp (rders_simp (RSTAR r0) (c#s)) =   (rders_simp (RSTAR r0) (c#s))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2199
  by (metis list.distinct(1) rders_simp_same_simpders star_closed_form1 star_closed_form2)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2200
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2201
lemma star_closed_form4:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2202
  shows " (rders_simp (RSTAR r0) (c#s)) = 
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2203
rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2204
  using star_closed_form2 star_closed_form3 by presburger
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2205
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2206
lemma star_closed_form5:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2207
  shows " rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0) ) Ss         )))) = 
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2208
          rsimp ( ( RALTS ( (map (\<lambda>s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss ))))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2209
  by (metis (mono_tags, lifting) list.map_comp map_eq_conv o_apply rsimp.simps(2) rsimp_idem)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2210
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2211
lemma star_closed_form6_hrewrites:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2212
  shows "  
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2213
 (map (\<lambda>s1.  (RSEQ (rsimp (rders r0 s1)) (RSTAR r0)) ) Ss )
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2214
 scf\<leadsto>*
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2215
(map (\<lambda>s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss )"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2216
  apply(induct Ss)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2217
  apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2218
  apply (simp add: ss1)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2219
  by (metis (no_types, lifting) list.simps(9) rsimp.simps(1) rsimp_idem simp_hrewrites ss2)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2220
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2221
lemma star_closed_form6:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2222
  shows " rsimp ( ( RALTS ( (map (\<lambda>s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss )))) = 
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2223
          rsimp ( ( RALTS ( (map (\<lambda>s1.  (RSEQ (rsimp (rders r0 s1)) (RSTAR r0)) ) Ss ))))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2224
  apply(subgoal_tac " map (\<lambda>s1.  (RSEQ (rsimp (rders r0 s1)) (RSTAR r0)) ) Ss  scf\<leadsto>*
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2225
                      map (\<lambda>s1.  rsimp (RSEQ  (rders r0 s1) (RSTAR r0)) ) Ss ")
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2226
  using hrewrites_simpeq srewritescf_alt1 apply fastforce
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2227
  using star_closed_form6_hrewrites by blast
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2228
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2229
lemma star_closed_form7:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2230
  shows  " (rders_simp (RSTAR r0) (c#s)) = 
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2231
rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rsimp (rders r0 s1)) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2232
  using star_closed_form4 star_closed_form5 star_closed_form6 by presburger
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2233
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2234
lemma derssimp_nonempty_list:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2235
  shows "\<forall>s \<in> set Ss. s \<noteq> [] \<Longrightarrow> 
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2236
 rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rsimp (rders r0 s1)) (RSTAR r0) ) Ss)))) =
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2237
 rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) Ss)))) "
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2238
  
487
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2239
  by (metis (no_types, lifting) map_eq_conv rders_simp_same_simpders)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2240
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2241
lemma stupdate_nonempty:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2242
  shows "\<forall>s \<in> set  Ss. s \<noteq> [] \<Longrightarrow> \<forall>s \<in> set (star_update c r Ss). s \<noteq> []"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2243
  apply(induct Ss)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2244
  apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2245
  apply(case_tac "rnullable (rders r a)")
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2246
   apply simp+
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2247
  done
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2248
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2249
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2250
lemma stupdates_nonempty:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2251
  shows "\<forall>s \<in> set Ss. s\<noteq> [] \<Longrightarrow> \<forall>s \<in> set (star_updates s r Ss). s \<noteq> []"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2252
  apply(induct s arbitrary: Ss)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2253
   apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2254
  apply simp
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2255
  using stupdate_nonempty by presburger
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2256
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2257
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2258
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2259
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2260
lemma star_closed_form8:
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2261
  shows  
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2262
"rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rsimp (rders r0 s1)) (RSTAR r0) ) (star_updates s r0 [[c]]) ) ))) = 
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2263
 rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ ( (rders_simp r0 s1)) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))"
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2264
  by (metis derssimp_nonempty_list neq_Nil_conv non_empty_list set_ConsD stupdates_nonempty)
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2265
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2266
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2267
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2268
lemma star_closed_form:
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2269
  shows "rders_simp (RSTAR r0) (c#s) = 
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2270
rsimp ( RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) ))"
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2271
  apply(induct s)
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2272
   apply simp
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2273
   apply (metis idem_after_simp1 rsimp.simps(1) rsimp.simps(6) rsimp_idem)
487
9f3d6f09b093 again starClosedForms
Chengsong
parents: 486
diff changeset
  2274
  using star_closed_form4 star_closed_form5 star_closed_form6 star_closed_form8 by presburger
486
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2275
f5b96a532c85 starclosed
Chengsong
parents: 485
diff changeset
  2276
476
Chengsong
parents: 475
diff changeset
  2277
444
a7e98deebb5c restructured sizebound proof
Chengsong
parents: 443
diff changeset
  2278
end