thys3/ClosedForms.thy
author Chengsong
Thu, 23 Jun 2022 16:09:40 +0100
changeset 543 b2bea5968b89
parent 513 ca7ca1f10f98
child 554 15d182ffbc76
permissions -rw-r--r--
thesis_thys
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
505
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
     1
theory ClosedForms 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
     2
  imports "BasicIdentities"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
     3
begin
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
     4
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
     5
lemma flts_middle0:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
     6
  shows "rflts (rsa @ RZERO # rsb) = rflts (rsa @ rsb)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
     7
  apply(induct rsa)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
     8
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
     9
  by (metis append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    10
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    11
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    12
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    13
lemma simp_flatten_aux0:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    14
  shows "rsimp (RALTS rs) = rsimp (RALTS (map rsimp rs))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    15
  by (metis append_Nil head_one_more_simp identity_wwo0 list.simps(8) rdistinct.simps(1) rflts.simps(1) rsimp.simps(2) rsimp_ALTs.simps(1) rsimp_ALTs.simps(3) simp_flatten spawn_simp_rsimpalts)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    16
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    17
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    18
inductive 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    19
  hrewrite:: "rrexp \<Rightarrow> rrexp \<Rightarrow> bool" ("_ h\<leadsto> _" [99, 99] 99)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    20
where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    21
  "RSEQ  RZERO r2 h\<leadsto> RZERO"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    22
| "RSEQ  r1 RZERO h\<leadsto> RZERO"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    23
| "RSEQ  RONE r h\<leadsto>  r"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    24
| "r1 h\<leadsto> r2 \<Longrightarrow> RSEQ  r1 r3 h\<leadsto> RSEQ r2 r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    25
| "r3 h\<leadsto> r4 \<Longrightarrow> RSEQ r1 r3 h\<leadsto> RSEQ r1 r4"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    26
| "r h\<leadsto> r' \<Longrightarrow> (RALTS (rs1 @ [r] @ rs2)) h\<leadsto> (RALTS  (rs1 @ [r'] @ rs2))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    27
(*context rule for eliminating 0, alts--corresponds to the recursive call flts r::rs = r::(flts rs)*)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    28
| "RALTS  (rsa @ [RZERO] @ rsb) h\<leadsto> RALTS  (rsa @ rsb)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    29
| "RALTS  (rsa @ [RALTS rs1] @ rsb) h\<leadsto> RALTS (rsa @ rs1 @ rsb)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    30
| "RALTS  [] h\<leadsto> RZERO"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    31
| "RALTS  [r] h\<leadsto> r"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    32
| "a1 = a2 \<Longrightarrow> RALTS (rsa@[a1]@rsb@[a2]@rsc) h\<leadsto> RALTS (rsa @ [a1] @ rsb @ rsc)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    33
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    34
inductive 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    35
  hrewrites:: "rrexp \<Rightarrow> rrexp \<Rightarrow> bool" ("_ h\<leadsto>* _" [100, 100] 100)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    36
where 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    37
  rs1[intro, simp]:"r h\<leadsto>* r"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    38
| rs2[intro]: "\<lbrakk>r1 h\<leadsto>* r2; r2 h\<leadsto> r3\<rbrakk> \<Longrightarrow> r1 h\<leadsto>* r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    39
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    40
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    41
lemma hr_in_rstar : "r1 h\<leadsto> r2 \<Longrightarrow> r1 h\<leadsto>* r2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    42
  using hrewrites.intros(1) hrewrites.intros(2) by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    43
 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    44
lemma hreal_trans[trans]: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    45
  assumes a1: "r1 h\<leadsto>* r2"  and a2: "r2 h\<leadsto>* r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    46
  shows "r1 h\<leadsto>* r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    47
  using a2 a1
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    48
  apply(induct r2 r3 arbitrary: r1 rule: hrewrites.induct) 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    49
  apply(auto)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    50
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    51
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    52
lemma hrewrites_seq_context:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    53
  shows "r1 h\<leadsto>* r2 \<Longrightarrow> RSEQ r1 r3 h\<leadsto>* RSEQ r2 r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    54
  apply(induct r1 r2 rule: hrewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    55
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    56
  using hrewrite.intros(4) by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    57
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    58
lemma hrewrites_seq_context2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    59
  shows "r1 h\<leadsto>* r2 \<Longrightarrow> RSEQ r0 r1 h\<leadsto>* RSEQ r0 r2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    60
  apply(induct r1 r2 rule: hrewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    61
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    62
  using hrewrite.intros(5) by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    63
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    64
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    65
lemma hrewrites_seq_contexts:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    66
  shows "\<lbrakk>r1 h\<leadsto>* r2; r3 h\<leadsto>* r4\<rbrakk> \<Longrightarrow> RSEQ r1 r3 h\<leadsto>* RSEQ r2 r4"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    67
  by (meson hreal_trans hrewrites_seq_context hrewrites_seq_context2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    68
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    69
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    70
lemma simp_removes_duplicate1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    71
  shows  " a \<in> set rsa \<Longrightarrow> rsimp (RALTS (rsa @ [a])) =  rsimp (RALTS (rsa))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    72
and " rsimp (RALTS (a1 # rsa @ [a1])) = rsimp (RALTS (a1 # rsa))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    73
  apply(induct rsa arbitrary: a1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    74
     apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    75
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    76
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    77
  apply(case_tac "a = aa")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    78
     apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    79
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    80
  apply (metis Cons_eq_appendI Cons_eq_map_conv distinct_removes_duplicate_flts list.set_intros(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    81
  apply (metis append_Cons append_Nil distinct_removes_duplicate_flts list.set_intros(1) list.simps(8) list.simps(9))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    82
  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))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    83
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    84
lemma simp_removes_duplicate2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    85
  shows "a \<in> set rsa \<Longrightarrow> rsimp (RALTS (rsa @ [a] @ rsb)) = rsimp (RALTS (rsa @ rsb))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    86
  apply(induct rsb arbitrary: rsa)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    87
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    88
  using distinct_removes_duplicate_flts apply auto[1]
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    89
  by (metis append.assoc head_one_more_simp rsimp.simps(2) simp_flatten simp_removes_duplicate1(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    90
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    91
lemma simp_removes_duplicate3:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    92
  shows "a \<in> set rsa \<Longrightarrow> rsimp (RALTS (rsa @ a # rsb)) = rsimp (RALTS (rsa @ rsb))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    93
  using simp_removes_duplicate2 by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    94
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    95
(*
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    96
lemma distinct_removes_middle4:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    97
  shows "a \<in> set rsa \<Longrightarrow> rdistinct (rsa @ [a] @ rsb) rset = rdistinct (rsa @ rsb) rset"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    98
  using distinct_removes_middle(1) by fastforce
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
    99
*)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   100
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   101
(*
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   102
lemma distinct_removes_middle_list:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   103
  shows "\<forall>a \<in> set x. a \<in> set rsa \<Longrightarrow> rdistinct (rsa @ x @ rsb) rset = rdistinct (rsa @ rsb) rset"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   104
  apply(induct x)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   105
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   106
  by (simp add: distinct_removes_middle3)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   107
*)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   108
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   109
inductive frewrite:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ \<leadsto>f _" [10, 10] 10)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   110
  where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   111
  "(RZERO # rs) \<leadsto>f rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   112
| "((RALTS rs) # rsa) \<leadsto>f (rs @ rsa)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   113
| "rs1 \<leadsto>f rs2 \<Longrightarrow> (r # rs1) \<leadsto>f (r # rs2)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   114
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   115
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   116
inductive 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   117
  frewrites:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ \<leadsto>f* _" [10, 10] 10)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   118
where 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   119
  [intro, simp]:"rs \<leadsto>f* rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   120
| [intro]: "\<lbrakk>rs1 \<leadsto>f* rs2; rs2 \<leadsto>f rs3\<rbrakk> \<Longrightarrow> rs1 \<leadsto>f* rs3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   121
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   122
inductive grewrite:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ \<leadsto>g _" [10, 10] 10)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   123
  where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   124
  "(RZERO # rs) \<leadsto>g rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   125
| "((RALTS rs) # rsa) \<leadsto>g (rs @ rsa)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   126
| "rs1 \<leadsto>g rs2 \<Longrightarrow> (r # rs1) \<leadsto>g (r # rs2)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   127
| "rsa @ [a] @ rsb @ [a] @ rsc \<leadsto>g rsa @ [a] @ rsb @ rsc" 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   128
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   129
lemma grewrite_variant1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   130
  shows "a \<in> set rs1 \<Longrightarrow> rs1 @ a # rs \<leadsto>g rs1 @ rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   131
  apply (metis append.assoc append_Cons append_Nil grewrite.intros(4) split_list_first)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   132
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   133
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   134
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   135
inductive 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   136
  grewrites:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ \<leadsto>g* _" [10, 10] 10)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   137
where 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   138
  [intro, simp]:"rs \<leadsto>g* rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   139
| [intro]: "\<lbrakk>rs1 \<leadsto>g* rs2; rs2 \<leadsto>g rs3\<rbrakk> \<Longrightarrow> rs1 \<leadsto>g* rs3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   140
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   141
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   142
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   143
(*
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   144
inductive 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   145
  frewrites2:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" ("_ <\<leadsto>f* _" [10, 10] 10)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   146
where 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   147
 [intro]: "\<lbrakk>rs1 \<leadsto>f* rs2; rs2 \<leadsto>f* rs1\<rbrakk> \<Longrightarrow> rs1 <\<leadsto>f* rs2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   148
*)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   149
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   150
lemma fr_in_rstar : "r1 \<leadsto>f r2 \<Longrightarrow> r1 \<leadsto>f* r2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   151
  using frewrites.intros(1) frewrites.intros(2) by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   152
 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   153
lemma freal_trans[trans]: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   154
  assumes a1: "r1 \<leadsto>f* r2"  and a2: "r2 \<leadsto>f* r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   155
  shows "r1 \<leadsto>f* r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   156
  using a2 a1
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   157
  apply(induct r2 r3 arbitrary: r1 rule: frewrites.induct) 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   158
  apply(auto)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   159
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   160
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   161
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   162
lemma  many_steps_later: "\<lbrakk>r1 \<leadsto>f r2; r2 \<leadsto>f* r3 \<rbrakk> \<Longrightarrow> r1 \<leadsto>f* r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   163
  by (meson fr_in_rstar freal_trans)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   164
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   165
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   166
lemma gr_in_rstar : "r1 \<leadsto>g r2 \<Longrightarrow> r1 \<leadsto>g* r2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   167
  using grewrites.intros(1) grewrites.intros(2) by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   168
 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   169
lemma greal_trans[trans]: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   170
  assumes a1: "r1 \<leadsto>g* r2"  and a2: "r2 \<leadsto>g* r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   171
  shows "r1 \<leadsto>g* r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   172
  using a2 a1
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   173
  apply(induct r2 r3 arbitrary: r1 rule: grewrites.induct) 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   174
  apply(auto)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   175
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   176
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   177
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   178
lemma  gmany_steps_later: "\<lbrakk>r1 \<leadsto>g r2; r2 \<leadsto>g* r3 \<rbrakk> \<Longrightarrow> r1 \<leadsto>g* r3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   179
  by (meson gr_in_rstar greal_trans)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   180
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   181
lemma gstar_rdistinct_general:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   182
  shows "rs1 @  rs \<leadsto>g* rs1 @ (rdistinct rs (set rs1))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   183
  apply(induct rs arbitrary: rs1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   184
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   185
  apply(case_tac " a \<in> set rs1")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   186
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   187
  apply(subgoal_tac "rs1 @ a # rs \<leadsto>g rs1 @ rs")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   188
  using gmany_steps_later apply auto[1]
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   189
  apply (metis append.assoc append_Cons append_Nil grewrite.intros(4) split_list_first)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   190
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   191
  apply(drule_tac x = "rs1 @ [a]" in meta_spec)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   192
  by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   193
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   194
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   195
lemma gstar_rdistinct:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   196
  shows "rs \<leadsto>g* rdistinct rs {}"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   197
  apply(induct rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   198
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   199
  by (metis append.left_neutral empty_set gstar_rdistinct_general)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   200
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   201
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   202
lemma grewrite_append:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   203
  shows "\<lbrakk> rsa \<leadsto>g rsb \<rbrakk> \<Longrightarrow> rs @ rsa \<leadsto>g rs @ rsb"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   204
  apply(induct rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   205
   apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   206
  using grewrite.intros(3) by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   207
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   208
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   209
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   210
lemma frewrites_cons:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   211
  shows "\<lbrakk> rsa \<leadsto>f* rsb \<rbrakk> \<Longrightarrow> r # rsa \<leadsto>f* r # rsb"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   212
  apply(induct rsa rsb rule: frewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   213
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   214
  using frewrite.intros(3) by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   215
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   216
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   217
lemma grewrites_cons:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   218
  shows "\<lbrakk> rsa \<leadsto>g* rsb \<rbrakk> \<Longrightarrow> r # rsa \<leadsto>g* r # rsb"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   219
  apply(induct rsa rsb rule: grewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   220
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   221
  using grewrite.intros(3) by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   222
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   223
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   224
lemma frewrites_append:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   225
  shows " \<lbrakk>rsa \<leadsto>f* rsb\<rbrakk> \<Longrightarrow> (rs @ rsa) \<leadsto>f* (rs @ rsb)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   226
  apply(induct rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   227
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   228
  by (simp add: frewrites_cons)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   229
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   230
lemma grewrites_append:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   231
  shows " \<lbrakk>rsa \<leadsto>g* rsb\<rbrakk> \<Longrightarrow> (rs @ rsa) \<leadsto>g* (rs @ rsb)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   232
  apply(induct rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   233
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   234
  by (simp add: grewrites_cons)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   235
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   236
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   237
lemma grewrites_concat:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   238
  shows "\<lbrakk>rs1 \<leadsto>g rs2; rsa \<leadsto>g* rsb \<rbrakk> \<Longrightarrow> (rs1 @ rsa) \<leadsto>g* (rs2 @ rsb)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   239
  apply(induct rs1 rs2 rule: grewrite.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   240
    apply(simp)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   241
  apply(subgoal_tac "(RZERO # rs @ rsa) \<leadsto>g (rs @ rsa)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   242
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   243
  using grewrite.intros(1) apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   244
    apply(subgoal_tac "(rs @ rsa) \<leadsto>g* (rs @ rsb)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   245
  using gmany_steps_later apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   246
  apply (simp add: grewrites_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   247
  apply (metis append.assoc append_Cons grewrite.intros(2) grewrites_append gmany_steps_later)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   248
  using grewrites_cons apply auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   249
  apply(subgoal_tac "rsaa @ a # rsba @ a # rsc @ rsa \<leadsto>g* rsaa @ a # rsba @ a # rsc @ rsb")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   250
  using grewrite.intros(4) grewrites.intros(2) apply force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   251
  using grewrites_append by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   252
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   253
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   254
lemma grewritess_concat:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   255
  shows "\<lbrakk>rsa \<leadsto>g* rsb; rsc \<leadsto>g* rsd \<rbrakk> \<Longrightarrow> (rsa @ rsc) \<leadsto>g* (rsb @ rsd)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   256
  apply(induct rsa rsb rule: grewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   257
   apply(case_tac rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   258
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   259
  using grewrites_append apply blast   
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   260
  by (meson greal_trans grewrites.simps grewrites_concat)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   261
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   262
fun alt_set:: "rrexp \<Rightarrow> rrexp set"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   263
  where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   264
  "alt_set (RALTS rs) = set rs \<union> \<Union> (alt_set ` (set rs))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   265
| "alt_set r = {r}"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   266
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   267
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   268
lemma grewrite_cases_middle:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   269
  shows "rs1 \<leadsto>g rs2 \<Longrightarrow> 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   270
(\<exists>rsa rsb rsc. rs1 =  (rsa @ [RALTS rsb] @ rsc) \<and> rs2 = (rsa @ rsb @ rsc)) \<or>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   271
(\<exists>rsa rsc. rs1 = rsa @ [RZERO] @ rsc \<and> rs2 = rsa @ rsc) \<or>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   272
(\<exists>rsa rsb rsc a. rs1 = rsa @ [a] @ rsb @ [a] @ rsc \<and> rs2 = rsa @ [a] @ rsb @ rsc)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   273
  apply( induct rs1 rs2 rule: grewrite.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   274
     apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   275
  apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   276
  apply (metis append_Cons append_Nil)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   277
  apply (metis append_Cons)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   278
  by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   279
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   280
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   281
lemma good_singleton:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   282
  shows "good a \<and> nonalt a  \<Longrightarrow> rflts [a] = [a]"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   283
  using good.simps(1) k0b by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   284
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   285
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   286
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   287
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   288
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   289
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   290
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   291
lemma all_that_same_elem:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   292
  shows "\<lbrakk> a \<in> rset; rdistinct rs {a} = []\<rbrakk>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   293
       \<Longrightarrow> rdistinct (rs @ rsb) rset = rdistinct rsb rset"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   294
  apply(induct rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   295
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   296
  apply(subgoal_tac "aa = a")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   297
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   298
  by (metis empty_iff insert_iff list.discI rdistinct.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   299
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   300
lemma distinct_early_app1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   301
  shows "rset1 \<subseteq> rset \<Longrightarrow> rdistinct rs rset = rdistinct (rdistinct rs rset1) rset"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   302
  apply(induct rs arbitrary: rset rset1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   303
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   304
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   305
  apply(case_tac "a \<in> rset1")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   306
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   307
   apply(case_tac "a \<in> rset")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   308
    apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   309
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   310
   apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   311
  apply(case_tac "a \<in> rset1")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   312
   apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   313
  apply(case_tac "a \<in> rset")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   314
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   315
   apply (metis insert_subsetI)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   316
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   317
  by (meson insert_mono)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   318
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   319
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   320
lemma distinct_early_app:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   321
  shows " rdistinct (rs @ rsb) rset = rdistinct (rdistinct rs {} @ rsb) rset"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   322
  apply(induct rsb)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   323
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   324
  using distinct_early_app1 apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   325
  by (metis distinct_early_app1 distinct_once_enough empty_subsetI)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   326
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   327
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   328
lemma distinct_eq_interesting1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   329
  shows "a \<in> rset \<Longrightarrow> rdistinct (rs @ rsb) rset = rdistinct (rdistinct (a # rs) {} @ rsb) rset"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   330
  apply(subgoal_tac "rdistinct (rdistinct (a # rs) {} @ rsb) rset = rdistinct (rdistinct rs {} @ rsb) rset")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   331
   apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   332
  using distinct_early_app apply blast 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   333
  by (metis append_Cons distinct_early_app rdistinct.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   334
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   335
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   336
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   337
lemma good_flatten_aux_aux1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   338
  shows "\<lbrakk> size rs \<ge>2; 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   339
\<forall>r \<in> set rs. good r \<and> r \<noteq> RZERO \<and> nonalt r; \<forall>r \<in> set rsb. good r \<and> r \<noteq> RZERO \<and> nonalt r \<rbrakk>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   340
       \<Longrightarrow> rdistinct (rs @ rsb) rset =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   341
           rdistinct (rflts [rsimp_ALTs (rdistinct rs {})] @ rsb) rset"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   342
  apply(induct rs arbitrary: rset)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   343
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   344
  apply(case_tac "a \<in> rset")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   345
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   346
   apply(case_tac "rdistinct rs {a}")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   347
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   348
    apply(subst good_singleton)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   349
     apply force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   350
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   351
    apply (meson all_that_same_elem)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   352
   apply(subgoal_tac "rflts [rsimp_ALTs (a # rdistinct rs {a})] = a # rdistinct rs {a} ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   353
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   354
  using k0a rsimp_ALTs.simps(3) apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   355
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   356
  apply(subgoal_tac "rdistinct (rs @ rsb) rset = rdistinct ((rdistinct (a # rs) {}) @ rsb) rset ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   357
    apply (metis insert_absorb insert_is_Un insert_not_empty rdistinct.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   358
   apply (meson distinct_eq_interesting1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   359
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   360
  apply(case_tac "rdistinct rs {a}")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   361
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   362
   apply(subgoal_tac "rsimp_ALTs (a # rdistinct rs {a}) = RALTS (a # rdistinct rs {a})")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   363
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   364
  apply(subgoal_tac "a # rdistinct (rs @ rsb) (insert a rset) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   365
           rdistinct (rflts [RALTS (a # rdistinct rs {a})] @ rsb) rset")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   366
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   367
  apply (metis append_Cons distinct_early_app empty_iff insert_is_Un k0a rdistinct.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   368
  using rsimp_ALTs.simps(3) apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   369
  by (metis Un_insert_left append_Cons distinct_early_app empty_iff good_singleton rdistinct.simps(2) rsimp_ALTs.simps(2) sup_bot_left)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   370
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   371
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   372
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   373
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   374
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   375
lemma good_flatten_aux_aux:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   376
  shows "\<lbrakk>\<exists>a aa lista list. rs = a # list \<and> list = aa # lista; 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   377
\<forall>r \<in> set rs. good r \<and> r \<noteq> RZERO \<and> nonalt r; \<forall>r \<in> set rsb. good r \<and> r \<noteq> RZERO \<and> nonalt r \<rbrakk>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   378
       \<Longrightarrow> rdistinct (rs @ rsb) rset =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   379
           rdistinct (rflts [rsimp_ALTs (rdistinct rs {})] @ rsb) rset"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   380
  apply(erule exE)+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   381
  apply(subgoal_tac "size rs \<ge> 2")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   382
   apply (metis good_flatten_aux_aux1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   383
  by (simp add: Suc_leI length_Cons less_add_Suc1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   384
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   385
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   386
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   387
lemma good_flatten_aux:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   388
  shows " \<lbrakk>\<forall>r\<in>set rs. good r \<or> r = RZERO; \<forall>r\<in>set rsa . good r \<or> r = RZERO; 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   389
           \<forall>r\<in>set rsb. good r \<or> r = RZERO;
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   390
     rsimp (RALTS (rsa @ rs @ rsb)) = rsimp_ALTs (rdistinct (rflts (rsa @ rs @ rsb)) {});
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   391
     rsimp (RALTS (rsa @ [RALTS rs] @ rsb)) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   392
     rsimp_ALTs (rdistinct (rflts (rsa @ [rsimp (RALTS rs)] @ rsb)) {});
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   393
     map rsimp rsa = rsa; map rsimp rsb = rsb; map rsimp rs = rs;
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   394
     rdistinct (rflts rsa @ rflts rs @ rflts rsb) {} =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   395
     rdistinct (rflts rsa) {} @ rdistinct (rflts rs @ rflts rsb) (set (rflts rsa));
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   396
     rdistinct (rflts rsa @ rflts [rsimp (RALTS rs)] @ rflts rsb) {} =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   397
     rdistinct (rflts rsa) {} @ rdistinct (rflts [rsimp (RALTS rs)] @ rflts rsb) (set (rflts rsa))\<rbrakk>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   398
    \<Longrightarrow>    rdistinct (rflts rs @ rflts rsb) rset =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   399
           rdistinct (rflts [rsimp (RALTS rs)] @ rflts rsb) rset"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   400
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   401
  apply(case_tac "rflts rs ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   402
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   403
  apply(case_tac "list")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   404
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   405
   apply(case_tac "a \<in> rset")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   406
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   407
  apply (metis append.left_neutral append_Cons equals0D k0b list.set_intros(1) nonalt_flts_rd qqq1 rdistinct.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   408
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   409
  apply (metis Un_insert_left append_Cons append_Nil ex_in_conv flts_single1 insertI1 list.simps(15) nonalt_flts_rd nonazero.elims(3) qqq1 rdistinct.simps(2) sup_bot_left)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   410
  apply(subgoal_tac "\<forall>r \<in> set (rflts rs). good r \<and> r \<noteq> RZERO \<and> nonalt r")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   411
   prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   412
  apply (metis Diff_empty flts3 nonalt_flts_rd qqq1 rdistinct_set_equality1)  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   413
  apply(subgoal_tac "\<forall>r \<in> set (rflts rsb). good r \<and> r \<noteq> RZERO \<and> nonalt r")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   414
   prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   415
  apply (metis Diff_empty flts3 good.simps(1) nonalt_flts_rd rdistinct_set_equality1)  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   416
  by (smt (verit, ccfv_threshold) good_flatten_aux_aux)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   417
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   418
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   419
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   420
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   421
lemma good_flatten_middle:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   422
  shows "\<lbrakk>\<forall>r \<in> set rs. good r \<or> r = RZERO; \<forall>r \<in> set rsa. good r \<or> r = RZERO; \<forall>r \<in> set rsb. good r \<or> r = RZERO\<rbrakk> \<Longrightarrow>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   423
rsimp (RALTS (rsa @ rs @ rsb)) = rsimp (RALTS (rsa @ [RALTS rs] @ rsb))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   424
  apply(subgoal_tac "rsimp (RALTS (rsa @ rs @ rsb)) = rsimp_ALTs (rdistinct (rflts (map rsimp rsa @ 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   425
map rsimp rs @ map rsimp rsb)) {})")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   426
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   427
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   428
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   429
    apply(subgoal_tac "rsimp (RALTS (rsa @ [RALTS rs] @ rsb)) = rsimp_ALTs (rdistinct (rflts (map rsimp rsa @ 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   430
[rsimp (RALTS rs)] @ map rsimp rsb)) {})")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   431
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   432
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   433
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   434
  apply(subgoal_tac "map rsimp rsa = rsa")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   435
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   436
  apply (metis map_idI rsimp.simps(3) test)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   437
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   438
  apply(subgoal_tac "map rsimp rsb = rsb")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   439
   prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   440
  apply (metis map_idI rsimp.simps(3) test)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   441
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   442
  apply(subst k00)+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   443
  apply(subgoal_tac "map rsimp rs = rs")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   444
   apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   445
   prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   446
  apply (metis map_idI rsimp.simps(3) test)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   447
  apply(subgoal_tac "rdistinct (rflts rsa @ rflts rs @ rflts rsb) {} = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   448
rdistinct (rflts rsa) {} @ rdistinct  (rflts rs @ rflts rsb) (set (rflts rsa))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   449
   apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   450
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   451
  using rdistinct_concat_general apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   452
  apply(subgoal_tac "rdistinct (rflts rsa @ rflts [rsimp (RALTS rs)] @ rflts rsb) {} = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   453
rdistinct (rflts rsa) {} @ rdistinct  (rflts [rsimp (RALTS rs)] @ rflts rsb) (set (rflts rsa))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   454
   apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   455
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   456
  using rdistinct_concat_general apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   457
  apply(subgoal_tac "rdistinct (rflts rs @ rflts rsb) (set (rflts rsa)) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   458
                     rdistinct  (rflts [rsimp (RALTS rs)] @ rflts rsb) (set (rflts rsa))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   459
   apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   460
  using good_flatten_aux by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   461
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   462
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   463
lemma simp_flatten3:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   464
  shows "rsimp (RALTS (rsa @ [RALTS rs] @ rsb)) = rsimp (RALTS (rsa @ rs @ rsb))"
511
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   465
proof -
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   466
  have  "rsimp (RALTS (rsa @ [RALTS rs] @ rsb)) = 
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   467
                     rsimp (RALTS (map rsimp rsa @ [rsimp (RALTS rs)] @ map rsimp rsb)) " 
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   468
    by (metis append_Cons append_Nil list.simps(9) map_append simp_flatten_aux0)
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   469
  apply(simp only:)
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   470
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   471
oops
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   472
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   473
lemma simp_flatten3:
47618d607bbf isarfied
Chengsong
parents: 506
diff changeset
   474
  shows "rsimp (RALTS (rsa @ [RALTS rs ] @ rsb)) = rsimp (RALTS (rsa @ rs @ rsb))"
505
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   475
  apply(subgoal_tac "rsimp (RALTS (rsa @ [RALTS rs] @ rsb)) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   476
                     rsimp (RALTS (map rsimp rsa @ [rsimp (RALTS rs)] @ map rsimp rsb)) ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   477
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   478
   apply (metis append.left_neutral append_Cons list.simps(9) map_append simp_flatten_aux0)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   479
  apply (simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   480
  apply(subgoal_tac "rsimp (RALTS (rsa @ rs @ rsb)) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   481
rsimp (RALTS (map rsimp rsa @ map rsimp rs @ map rsimp rsb))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   482
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   483
   apply (metis map_append simp_flatten_aux0)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   484
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   485
  apply(subgoal_tac "rsimp  (RALTS (map rsimp rsa @ map rsimp rs @ map rsimp rsb)) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   486
 rsimp (RALTS (map rsimp rsa @ [RALTS (map rsimp rs)] @ map rsimp rsb))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   487
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   488
   apply (metis (no_types, lifting) head_one_more_simp map_append simp_flatten_aux0)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   489
  apply(subgoal_tac "\<forall>r \<in> set (map rsimp rsa). good r \<or> r = RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   490
   apply(subgoal_tac "\<forall>r \<in> set (map rsimp rs). good r \<or> r = RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   491
    apply(subgoal_tac "\<forall>r \<in> set (map rsimp rsb). good r \<or> r = RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   492
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   493
  using good_flatten_middle apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   494
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   495
  apply (simp add: good1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   496
  apply (simp add: good1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   497
  apply (simp add: good1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   498
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   499
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   500
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   501
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   502
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   503
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   504
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   505
lemma grewrite_equal_rsimp:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   506
  shows "rs1 \<leadsto>g rs2 \<Longrightarrow> rsimp (RALTS rs1) = rsimp (RALTS rs2)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   507
  apply(frule grewrite_cases_middle)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   508
  apply(case_tac "(\<exists>rsa rsb rsc. rs1 = rsa @ [RALTS rsb] @ rsc \<and> rs2 = rsa @ rsb @ rsc)")  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   509
  using simp_flatten3 apply auto[1]
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   510
  apply(case_tac "(\<exists>rsa rsc. rs1 = rsa @ [RZERO] @ rsc \<and> rs2 = rsa @ rsc)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   511
  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)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   512
  by (smt (verit) append.assoc append_Cons append_Nil in_set_conv_decomp simp_removes_duplicate3)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   513
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   514
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   515
lemma grewrites_equal_rsimp:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   516
  shows "rs1 \<leadsto>g* rs2 \<Longrightarrow> rsimp (RALTS rs1) = rsimp (RALTS rs2)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   517
  apply (induct rs1 rs2 rule: grewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   518
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   519
  using grewrite_equal_rsimp by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   520
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   521
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   522
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   523
lemma grewrites_last:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   524
  shows "r # [RALTS rs] \<leadsto>g*  r # rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   525
  by (metis gr_in_rstar grewrite.intros(2) grewrite.intros(3) self_append_conv)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   526
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   527
lemma simp_flatten2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   528
  shows "rsimp (RALTS (r # [RALTS rs])) = rsimp (RALTS (r # rs))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   529
  using grewrites_equal_rsimp grewrites_last by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   530
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   531
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   532
lemma frewrites_alt:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   533
  shows "rs1 \<leadsto>f* rs2 \<Longrightarrow> (RALT r1 r2) # rs1 \<leadsto>f* r1 # r2 # rs2"  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   534
  by (metis Cons_eq_appendI append_self_conv2 frewrite.intros(2) frewrites_cons many_steps_later)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   535
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   536
lemma early_late_der_frewrites:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   537
  shows "map (rder x) (rflts rs) \<leadsto>f* rflts (map (rder x) rs)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   538
  apply(induct rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   539
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   540
  apply(case_tac a)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   541
       apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   542
  using frewrite.intros(1) many_steps_later apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   543
     apply(case_tac "x = x3")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   544
      apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   545
  using frewrites_cons apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   546
  using frewrite.intros(1) many_steps_later apply fastforce
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   547
  apply(case_tac "rnullable x41")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   548
     apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   549
     apply (simp add: frewrites_alt)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   550
  apply (simp add: frewrites_cons)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   551
   apply (simp add: frewrites_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   552
  by (simp add: frewrites_cons)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   553
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   554
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   555
lemma gstar0:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   556
  shows "rsa @ (rdistinct rs (set rsa)) \<leadsto>g* rsa @ (rdistinct rs (insert RZERO (set rsa)))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   557
  apply(induct rs arbitrary: rsa)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   558
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   559
  apply(case_tac "a = RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   560
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   561
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   562
  using gr_in_rstar grewrite.intros(1) grewrites_append apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   563
  apply(case_tac "a \<in> set rsa")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   564
   apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   565
  apply(drule_tac x = "rsa @ [a]" in meta_spec)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   566
  by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   567
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   568
lemma grewrite_rdistinct_aux:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   569
  shows "rs @ rdistinct rsa rset \<leadsto>g* rs @ rdistinct rsa (rset \<union> set rs)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   570
  apply(induct rsa arbitrary: rs rset)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   571
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   572
  apply(case_tac " a \<in> rset")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   573
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   574
  apply(case_tac "a \<in> set rs")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   575
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   576
   apply (metis Un_insert_left Un_insert_right gmany_steps_later grewrite_variant1 insert_absorb)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   577
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   578
  apply(drule_tac x = "rs @ [a]" in meta_spec)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   579
  by (metis Un_insert_left Un_insert_right append.assoc append.right_neutral append_Cons append_Nil insert_absorb2 list.simps(15) set_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   580
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   581
 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   582
lemma flts_gstar:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   583
  shows "rs \<leadsto>g* rflts rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   584
  apply(induct rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   585
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   586
  apply(case_tac "a = RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   587
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   588
  using gmany_steps_later grewrite.intros(1) apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   589
  apply(case_tac "\<exists>rsa. a = RALTS rsa")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   590
   apply(erule exE)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   591
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   592
   apply (meson grewrite.intros(2) grewrites.simps grewrites_cons)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   593
  by (simp add: grewrites_cons rflts_def_idiot)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   594
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   595
lemma more_distinct1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   596
  shows "       \<lbrakk>\<And>rsb rset rset2.
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   597
           rset2 \<subseteq> set rsb \<Longrightarrow> rsb @ rdistinct rs rset \<leadsto>g* rsb @ rdistinct rs (rset \<union> rset2);
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   598
        rset2 \<subseteq> set rsb; a \<notin> rset; a \<in> rset2\<rbrakk>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   599
       \<Longrightarrow> rsb @ a # rdistinct rs (insert a rset) \<leadsto>g* rsb @ rdistinct rs (rset \<union> rset2)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   600
  apply(subgoal_tac "rsb @ a # rdistinct rs (insert a rset) \<leadsto>g* rsb @ rdistinct rs (insert a rset)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   601
   apply(subgoal_tac "rsb @ rdistinct rs (insert a rset) \<leadsto>g* rsb @ rdistinct rs (rset \<union> rset2)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   602
    apply (meson greal_trans)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   603
   apply (metis Un_iff Un_insert_left insert_absorb)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   604
  by (simp add: gr_in_rstar grewrite_variant1 in_mono)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   605
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   606
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   607
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   608
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   609
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   610
lemma frewrite_rd_grewrites_aux:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   611
  shows     "       RALTS rs \<notin> set rsb \<Longrightarrow>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   612
       rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   613
       RALTS rs #
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   614
       rdistinct rsa
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   615
        (insert (RALTS rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   616
          (set rsb)) \<leadsto>g* rflts rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   617
                          rdistinct rs (set rsb) @ rdistinct rsa (set rs \<union> set rsb \<union> {RALTS rs})"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   618
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   619
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   620
  apply(subgoal_tac "rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   621
    RALTS rs #
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   622
    rdistinct rsa
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   623
     (insert (RALTS rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   624
       (set rsb)) \<leadsto>g* rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   625
    rs @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   626
    rdistinct rsa
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   627
     (insert (RALTS rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   628
       (set rsb)) ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   629
  apply(subgoal_tac " rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   630
    rs @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   631
    rdistinct rsa
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   632
     (insert (RALTS rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   633
       (set rsb)) \<leadsto>g*
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   634
                      rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   635
    rdistinct rs (set rsb) @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   636
    rdistinct rsa
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   637
     (insert (RALTS rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   638
       (set rsb)) ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   639
    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)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   640
   apply (metis append_assoc grewrites.intros(1) grewritess_concat gstar_rdistinct_general)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   641
  by (simp add: gr_in_rstar grewrite.intros(2) grewrites_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   642
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   643
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   644
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   645
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   646
lemma list_dlist_union:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   647
  shows "set rs \<subseteq> set rsb \<union> set (rdistinct rs (set rsb))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   648
  by (metis rdistinct_concat_general rdistinct_set_equality set_append sup_ge2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   649
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   650
lemma r_finite1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   651
  shows "r = RALTS (r # rs) = False"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   652
  apply(induct r)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   653
  apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   654
   apply (metis list.set_intros(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   655
  by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   656
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   657
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   658
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   659
lemma grewrite_singleton:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   660
  shows "[r] \<leadsto>g r # rs \<Longrightarrow> rs = []"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   661
  apply (induct "[r]" "r # rs" rule: grewrite.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   662
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   663
  apply (metis r_finite1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   664
  using grewrite.simps apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   665
  by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   666
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   667
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   668
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   669
lemma concat_rdistinct_equality1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   670
  shows "rdistinct (rs @ rsa) rset = rdistinct rs rset @ rdistinct rsa (rset \<union> (set rs))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   671
  apply(induct rs arbitrary: rsa rset)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   672
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   673
  apply(case_tac "a \<in> rset")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   674
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   675
  apply (simp add: insert_absorb)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   676
  by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   677
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   678
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   679
lemma grewrites_rev_append:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   680
  shows "rs1 \<leadsto>g* rs2 \<Longrightarrow> rs1 @ [x] \<leadsto>g* rs2 @ [x]"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   681
  using grewritess_concat by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   682
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   683
lemma grewrites_inclusion:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   684
  shows "set rs \<subseteq> set rs1 \<Longrightarrow> rs1 @ rs \<leadsto>g* rs1"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   685
  apply(induct rs arbitrary: rs1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   686
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   687
  by (meson gmany_steps_later grewrite_variant1 list.set_intros(1) set_subset_Cons subset_code(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   688
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   689
lemma distinct_keeps_last:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   690
  shows "\<lbrakk>x \<notin> rset; x \<notin> set xs \<rbrakk> \<Longrightarrow> rdistinct (xs @ [x]) rset = rdistinct xs rset @ [x]"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   691
  by (simp add: concat_rdistinct_equality1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   692
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   693
lemma grewrites_shape2_aux:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   694
  shows "       RALTS rs \<notin> set rsb \<Longrightarrow>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   695
       rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   696
       rdistinct (rs @ rsa)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   697
        (set rsb) \<leadsto>g* rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   698
                       rdistinct rs (set rsb) @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   699
                       rdistinct rsa (set rs \<union> set rsb \<union> {RALTS rs})"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   700
  apply(subgoal_tac " rdistinct (rs @ rsa) (set rsb) =  rdistinct rs (set rsb) @ rdistinct rsa (set rs \<union> set rsb)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   701
   apply (simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   702
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   703
  apply (simp add: Un_commute concat_rdistinct_equality1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   704
  apply(induct rsa arbitrary: rs rsb rule: rev_induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   705
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   706
  apply(case_tac "x \<in> set rs")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   707
  apply (simp add: distinct_removes_middle3)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   708
  apply(case_tac "x = RALTS rs")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   709
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   710
  apply(case_tac "x \<in> set rsb")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   711
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   712
    apply (simp add: concat_rdistinct_equality1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   713
  apply (simp add: concat_rdistinct_equality1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   714
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   715
  apply(drule_tac x = "rs " in meta_spec)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   716
  apply(drule_tac x = rsb in meta_spec)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   717
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   718
  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))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   719
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   720
   apply (simp add: concat_rdistinct_equality1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   721
  apply(case_tac "x \<in> set xs")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   722
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   723
   apply (simp add: distinct_removes_last)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   724
  apply(case_tac "x \<in> set rsb")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   725
   apply (smt (verit, ccfv_threshold) Un_iff append.right_neutral concat_rdistinct_equality1 insert_is_Un rdistinct.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   726
  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]")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   727
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   728
  apply(case_tac "x = RALTS rs")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   729
    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")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   730
  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) ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   731
      apply (smt (verit, ccfv_SIG) Un_insert_left append.right_neutral concat_rdistinct_equality1 greal_trans insert_iff rdistinct.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   732
  apply(subgoal_tac "set rs \<subseteq> set ( rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \<union> set rsb))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   733
  apply (metis append.assoc grewrites_inclusion)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   734
     apply (metis Un_upper1 append.assoc dual_order.trans list_dlist_union set_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   735
  apply (metis append_Nil2 gr_in_rstar grewrite.intros(2) grewrite_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   736
   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]")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   737
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   738
  apply (metis append.assoc grewrites_rev_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   739
  apply (simp add: insert_absorb)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   740
   apply (simp add: distinct_keeps_last)+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   741
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   742
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   743
lemma grewrites_shape2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   744
  shows "       RALTS rs \<notin> set rsb \<Longrightarrow>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   745
       rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   746
       rdistinct (rs @ rsa)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   747
        (set rsb) \<leadsto>g* rflts rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   748
                       rdistinct rs (set rsb) @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   749
                       rdistinct rsa (set rs \<union> set rsb \<union> {RALTS rs})"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   750
  apply (meson flts_gstar greal_trans grewrites.simps grewrites_shape2_aux grewritess_concat)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   751
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   752
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   753
lemma rdistinct_add_acc:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   754
  shows "rset2 \<subseteq> set rsb \<Longrightarrow> rsb @ rdistinct rs rset \<leadsto>g* rsb @ rdistinct rs (rset \<union> rset2)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   755
  apply(induct rs arbitrary: rsb rset rset2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   756
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   757
  apply (case_tac "a \<in> rset")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   758
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   759
  apply(case_tac "a \<in> rset2")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   760
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   761
  apply (simp add: more_distinct1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   762
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   763
  apply(drule_tac x = "rsb @ [a]" in meta_spec)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   764
  by (metis Un_insert_left append.assoc append_Cons append_Nil set_append sup.coboundedI1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   765
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   766
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   767
lemma frewrite_fun1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   768
  shows "       RALTS rs \<in> set rsb \<Longrightarrow>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   769
       rsb @ rdistinct rsa (set rsb) \<leadsto>g* rflts rsb @ rdistinct rsa (set rsb \<union> set rs)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   770
  apply(subgoal_tac "rsb @ rdistinct rsa (set rsb) \<leadsto>g* rflts rsb @ rdistinct rsa (set rsb)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   771
   apply(subgoal_tac " set rs \<subseteq> set (rflts rsb)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   772
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   773
  using spilled_alts_contained apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   774
   apply(subgoal_tac "rflts rsb @ rdistinct rsa (set rsb) \<leadsto>g* rflts rsb @ rdistinct rsa (set rsb \<union> set rs)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   775
  using greal_trans apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   776
  using rdistinct_add_acc apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   777
  using flts_gstar grewritess_concat by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   778
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   779
lemma frewrite_rd_grewrites:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   780
  shows "rs1 \<leadsto>f rs2 \<Longrightarrow> 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   781
\<exists>rs3. (rs @ (rdistinct rs1 (set rs)) \<leadsto>g* rs3) \<and> (rs @ (rdistinct rs2 (set rs)) \<leadsto>g* rs3) "
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   782
  apply(induct rs1 rs2 arbitrary: rs rule: frewrite.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   783
    apply(rule_tac x = "rsa @ (rdistinct rs ({RZERO} \<union> set rsa))" in exI)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   784
    apply(rule conjI)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   785
  apply(case_tac "RZERO \<in> set rsa")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   786
  apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   787
  using gstar0 apply fastforce
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   788
     apply (simp add: gr_in_rstar grewrite.intros(1) grewrites_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   789
    apply (simp add: gstar0)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   790
    prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   791
    apply(case_tac "r \<in> set rs")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   792
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   793
    apply(drule_tac x = "rs @ [r]" in meta_spec)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   794
    apply(erule exE)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   795
    apply(rule_tac x = "rs3" in exI)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   796
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   797
  apply(case_tac "RALTS rs \<in> set rsb")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   798
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   799
   apply(rule_tac x = "rflts rsb @ rdistinct rsa (set rsb \<union> set rs)" in exI)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   800
   apply(rule conjI)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   801
  using frewrite_fun1 apply force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   802
  apply (metis frewrite_fun1 rdistinct_concat sup_ge2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   803
  apply(simp)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   804
  apply(rule_tac x = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   805
 "rflts rsb @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   806
                       rdistinct rs (set rsb) @
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   807
                       rdistinct rsa (set rs \<union> set rsb \<union> {RALTS rs})" in exI)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   808
  apply(rule conjI)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   809
   prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   810
  using grewrites_shape2 apply force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   811
  using frewrite_rd_grewrites_aux by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   812
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   813
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   814
lemma frewrite_simpeq2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   815
  shows "rs1 \<leadsto>f rs2 \<Longrightarrow> rsimp (RALTS (rdistinct rs1 {})) = rsimp (RALTS (rdistinct rs2 {}))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   816
  apply(subgoal_tac "\<exists> rs3. (rdistinct rs1 {} \<leadsto>g* rs3) \<and> (rdistinct rs2 {} \<leadsto>g* rs3)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   817
  using grewrites_equal_rsimp apply fastforce
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   818
  by (metis append_self_conv2 frewrite_rd_grewrites list.set(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   819
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   820
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   821
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   822
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   823
(*a more refined notion of h\<leadsto>* is needed,
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   824
this lemma fails when rs1 contains some RALTS rs where elements
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   825
of rs appear in later parts of rs1, which will be picked up by rs2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   826
and deduplicated*)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   827
lemma frewrites_simpeq:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   828
  shows "rs1 \<leadsto>f* rs2 \<Longrightarrow>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   829
 rsimp (RALTS (rdistinct rs1 {})) = rsimp (RALTS ( rdistinct rs2 {})) "
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   830
  apply(induct rs1 rs2 rule: frewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   831
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   832
  using frewrite_simpeq2 by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   833
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   834
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   835
lemma frewrite_single_step:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   836
  shows " rs2 \<leadsto>f rs3 \<Longrightarrow> rsimp (RALTS rs2) = rsimp (RALTS rs3)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   837
  apply(induct rs2 rs3 rule: frewrite.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   838
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   839
  using simp_flatten apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   840
  by (metis (no_types, opaque_lifting) list.simps(9) rsimp.simps(2) simp_flatten2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   841
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   842
lemma grewrite_simpalts:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   843
  shows " rs2 \<leadsto>g rs3 \<Longrightarrow> rsimp (rsimp_ALTs rs2) = rsimp (rsimp_ALTs rs3)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   844
  apply(induct rs2 rs3 rule : grewrite.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   845
  using identity_wwo0 apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   846
  apply (metis frewrite.intros(1) frewrite_single_step identity_wwo0 rsimp_ALTs.simps(3) simp_flatten)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   847
  apply (smt (verit, ccfv_SIG) gmany_steps_later grewrites.intros(1) grewrites_cons grewrites_equal_rsimp identity_wwo0 rsimp_ALTs.simps(3))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   848
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   849
  apply(subst rsimp_alts_equal)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   850
  apply(case_tac "rsa = [] \<and> rsb = [] \<and> rsc = []")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   851
   apply(subgoal_tac "rsa @ a # rsb @ rsc = [a]")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   852
  apply (simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   853
  apply (metis append_Nil frewrite.intros(1) frewrite_single_step identity_wwo0 rsimp_ALTs.simps(3) simp_removes_duplicate1(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   854
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   855
  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)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   856
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   857
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   858
lemma grewrites_simpalts:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   859
  shows " rs2 \<leadsto>g* rs3 \<Longrightarrow> rsimp (rsimp_ALTs rs2) = rsimp (rsimp_ALTs rs3)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   860
  apply(induct rs2 rs3 rule: grewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   861
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   862
  using grewrite_simpalts by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   863
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   864
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   865
lemma simp_der_flts:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   866
  shows "rsimp (RALTS (rdistinct (map (rder x) (rflts rs)) {})) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   867
         rsimp (RALTS (rdistinct (rflts (map (rder x) rs)) {}))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   868
  apply(subgoal_tac "map (rder x) (rflts rs) \<leadsto>f* rflts (map (rder x) rs)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   869
  using frewrites_simpeq apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   870
  using early_late_der_frewrites by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   871
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   872
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   873
lemma simp_der_pierce_flts_prelim:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   874
  shows "rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts rs)) {})) 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   875
       = rsimp (rsimp_ALTs (rdistinct (rflts (map (rder x) rs)) {}))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   876
  by (metis append.right_neutral grewrite.intros(2) grewrite_simpalts rsimp_ALTs.simps(2) simp_der_flts)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   877
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   878
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   879
lemma basic_regex_property1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   880
  shows "rnullable r \<Longrightarrow> rsimp r \<noteq> RZERO"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   881
  apply(induct r rule: rsimp.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   882
  apply(auto)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   883
  apply (metis idiot idiot2 rrexp.distinct(5))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   884
  by (metis der_simp_nullability rnullable.simps(1) rnullable.simps(4) rsimp.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   885
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   886
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   887
lemma inside_simp_seq_nullable:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   888
  shows 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   889
"\<And>r1 r2.
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   890
       \<lbrakk>rsimp (rder x (rsimp r1)) = rsimp (rder x r1); rsimp (rder x (rsimp r2)) = rsimp (rder x r2);
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   891
        rnullable r1\<rbrakk>
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   892
       \<Longrightarrow> rsimp (rder x (rsimp_SEQ (rsimp r1) (rsimp r2))) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   893
           rsimp_ALTs (rdistinct (rflts [rsimp_SEQ (rsimp (rder x r1)) (rsimp r2), rsimp (rder x r2)]) {})"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   894
  apply(case_tac "rsimp r1 = RONE")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   895
   apply(simp)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   896
  apply(subst basic_rsimp_SEQ_property1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   897
   apply (simp add: idem_after_simp1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   898
  apply(case_tac "rsimp r1 = RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   899
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   900
  using basic_regex_property1 apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   901
  apply(case_tac "rsimp r2 = RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   902
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   903
  apply (simp add: basic_rsimp_SEQ_property3)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   904
  apply(subst idiot2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   905
     apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   906
  apply(subgoal_tac "rnullable (rsimp r1)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   907
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   908
  using rsimp_idem apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   909
  using der_simp_nullability by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   910
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   911
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   912
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   913
lemma grewrite_ralts:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   914
  shows "rs \<leadsto>g rs' \<Longrightarrow> RALTS rs h\<leadsto>* RALTS rs'"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   915
  by (smt (verit) grewrite_cases_middle hr_in_rstar hrewrite.intros(11) hrewrite.intros(7) hrewrite.intros(8))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   916
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   917
lemma grewrites_ralts:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   918
  shows "rs \<leadsto>g* rs' \<Longrightarrow> RALTS rs h\<leadsto>* RALTS rs'"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   919
  apply(induct rule: grewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   920
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   921
  using grewrite_ralts hreal_trans by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   922
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   923
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   924
lemma distinct_grewrites_subgoal1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   925
  shows "  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   926
       \<lbrakk>rs1 \<leadsto>g* [a]; RALTS rs1 h\<leadsto>* a; [a] \<leadsto>g rs3\<rbrakk> \<Longrightarrow> RALTS rs1 h\<leadsto>* rsimp_ALTs rs3"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   927
  apply(subgoal_tac "RALTS rs1 h\<leadsto>* RALTS rs3")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   928
  apply (metis hrewrite.intros(10) hrewrite.intros(9) rs2 rsimp_ALTs.cases rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   929
  apply(subgoal_tac "rs1 \<leadsto>g* rs3")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   930
  using grewrites_ralts apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   931
  using grewrites.intros(2) by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   932
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   933
lemma grewrites_ralts_rsimpalts:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   934
  shows "rs \<leadsto>g* rs' \<Longrightarrow> RALTS rs h\<leadsto>* rsimp_ALTs rs' "
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   935
  apply(induct rs rs' rule: grewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   936
   apply(case_tac rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   937
  using hrewrite.intros(9) apply force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   938
   apply(case_tac list)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   939
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   940
  using hr_in_rstar hrewrite.intros(10) rsimp_ALTs.simps(2) apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   941
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   942
  apply(case_tac rs2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   943
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   944
   apply (metis grewrite.intros(3) grewrite_singleton rsimp_ALTs.simps(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   945
  apply(case_tac list)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   946
   apply(simp)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   947
  using distinct_grewrites_subgoal1 apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   948
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   949
  apply(case_tac rs3)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   950
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   951
  using grewrites_ralts hrewrite.intros(9) apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   952
  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))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   953
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   954
lemma hrewrites_alts:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   955
  shows " r h\<leadsto>* r' \<Longrightarrow> (RALTS (rs1 @ [r] @ rs2)) h\<leadsto>* (RALTS  (rs1 @ [r'] @ rs2))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   956
  apply(induct r r' rule: hrewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   957
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   958
  using hrewrite.intros(6) by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   959
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   960
inductive 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   961
  srewritescf:: "rrexp list \<Rightarrow> rrexp list \<Rightarrow> bool" (" _ scf\<leadsto>* _" [100, 100] 100)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   962
where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   963
  ss1: "[] scf\<leadsto>* []"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   964
| ss2: "\<lbrakk>r h\<leadsto>* r'; rs scf\<leadsto>* rs'\<rbrakk> \<Longrightarrow> (r#rs) scf\<leadsto>* (r'#rs')"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   965
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   966
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   967
lemma srewritescf_alt: "rs1 scf\<leadsto>* rs2 \<Longrightarrow> (RALTS (rs@rs1)) h\<leadsto>* (RALTS (rs@rs2))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   968
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   969
  apply(induct rs1 rs2 arbitrary: rs rule: srewritescf.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   970
   apply(rule rs1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   971
  apply(drule_tac x = "rsa@[r']" in meta_spec)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   972
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   973
  apply(rule hreal_trans)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   974
   prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   975
   apply(assumption)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   976
  apply(drule hrewrites_alts)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   977
  by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   978
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   979
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   980
corollary srewritescf_alt1: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   981
  assumes "rs1 scf\<leadsto>* rs2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   982
  shows "RALTS rs1 h\<leadsto>* RALTS rs2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   983
  using assms
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   984
  by (metis append_Nil srewritescf_alt)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   985
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   986
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   987
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   988
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   989
lemma trivialrsimp_srewrites: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   990
  "\<lbrakk>\<And>x. x \<in> set rs \<Longrightarrow> x h\<leadsto>* f x \<rbrakk> \<Longrightarrow> rs scf\<leadsto>* (map f rs)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   991
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   992
  apply(induction rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   993
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   994
   apply(rule ss1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   995
  by (metis insert_iff list.simps(15) list.simps(9) srewritescf.simps)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   996
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   997
lemma hrewrites_list: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   998
  shows
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
   999
" (\<And>xa. xa \<in> set x \<Longrightarrow> xa h\<leadsto>* rsimp xa) \<Longrightarrow> RALTS x h\<leadsto>* RALTS (map rsimp x)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1000
  apply(induct x)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1001
   apply(simp)+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1002
  by (simp add: srewritescf_alt1 ss2 trivialrsimp_srewrites)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1003
(*  apply(subgoal_tac "RALTS x h\<leadsto>* RALTS (map rsimp x)")*)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1004
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1005
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1006
lemma hrewrite_simpeq:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1007
  shows "r1 h\<leadsto> r2 \<Longrightarrow> rsimp r1 = rsimp r2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1008
  apply(induct rule: hrewrite.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1009
            apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1010
  apply (simp add: basic_rsimp_SEQ_property3)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1011
  apply (simp add: basic_rsimp_SEQ_property1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1012
  using rsimp.simps(1) apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1013
        apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1014
  using flts_middle0 apply force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1015
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1016
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1017
  using simp_flatten3 apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1018
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1019
  apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1020
  apply (simp add: idem_after_simp1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1021
  using grewrite.intros(4) grewrite_equal_rsimp by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1022
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1023
lemma hrewrites_simpeq:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1024
  shows "r1 h\<leadsto>* r2 \<Longrightarrow> rsimp r1 = rsimp r2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1025
  apply(induct rule: hrewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1026
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1027
  apply(subgoal_tac "rsimp r2 = rsimp r3")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1028
   apply auto[1]
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1029
  using hrewrite_simpeq by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1030
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1031
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1032
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1033
lemma simp_hrewrites:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1034
  shows "r1 h\<leadsto>* rsimp r1"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1035
  apply(induct r1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1036
       apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1037
    apply(case_tac "rsimp r11 = RONE")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1038
     apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1039
     apply(subst basic_rsimp_SEQ_property1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1040
  apply(subgoal_tac "RSEQ r11 r12 h\<leadsto>* RSEQ RONE r12")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1041
  using hreal_trans hrewrite.intros(3) apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1042
  using hrewrites_seq_context apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1043
    apply(case_tac "rsimp r11 = RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1044
     apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1045
  using hrewrite.intros(1) hrewrites_seq_context apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1046
    apply(case_tac "rsimp r12 = RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1047
     apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1048
  apply(subst basic_rsimp_SEQ_property3)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1049
  apply (meson hrewrite.intros(2) hrewrites.simps hrewrites_seq_context2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1050
    apply(subst idiot2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1051
       apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1052
  using hrewrites_seq_contexts apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1053
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1054
   apply(subgoal_tac "RALTS x h\<leadsto>* RALTS (map rsimp x)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1055
  apply(subgoal_tac "RALTS (map rsimp x) h\<leadsto>* rsimp_ALTs (rdistinct (rflts (map rsimp x)) {}) ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1056
  using hreal_trans apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1057
    apply (meson flts_gstar greal_trans grewrites_ralts_rsimpalts gstar_rdistinct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1058
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1059
   apply (simp add: grewrites_ralts hrewrites_list)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1060
  by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1061
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1062
lemma interleave_aux1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1063
  shows " RALT (RSEQ RZERO r1) r h\<leadsto>*  r"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1064
  apply(subgoal_tac "RSEQ RZERO r1 h\<leadsto>* RZERO")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1065
  apply(subgoal_tac "RALT (RSEQ RZERO r1) r h\<leadsto>* RALT RZERO r")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1066
  apply (meson grewrite.intros(1) grewrite_ralts hreal_trans hrewrite.intros(10) hrewrites.simps)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1067
  using rs1 srewritescf_alt1 ss1 ss2 apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1068
  by (simp add: hr_in_rstar hrewrite.intros(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1069
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1070
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1071
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1072
lemma rnullable_hrewrite:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1073
  shows "r1 h\<leadsto> r2 \<Longrightarrow> rnullable r1 = rnullable r2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1074
  apply(induct rule: hrewrite.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1075
            apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1076
     apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1077
    apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1078
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1079
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1080
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1081
lemma interleave1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1082
  shows "r h\<leadsto> r' \<Longrightarrow> rder c r h\<leadsto>* rder c r'"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1083
  apply(induct r r' rule: hrewrite.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1084
            apply (simp add: hr_in_rstar hrewrite.intros(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1085
  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)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1086
          apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1087
          apply(subst interleave_aux1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1088
          apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1089
         apply(case_tac "rnullable r1")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1090
          apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1091
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1092
          apply (simp add: hrewrites_seq_context rnullable_hrewrite srewritescf_alt1 ss1 ss2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1093
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1094
         apply (simp add: hrewrites_seq_context rnullable_hrewrite)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1095
        apply(case_tac "rnullable r1")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1096
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1097
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1098
  using hr_in_rstar hrewrites_seq_context2 srewritescf_alt1 ss1 ss2 apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1099
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1100
  using hr_in_rstar hrewrites_seq_context2 apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1101
       apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1102
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1103
  using hrewrites_alts apply auto[1]
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1104
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1105
  using grewrite.intros(1) grewrite_append grewrite_ralts apply auto[1]
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1106
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1107
  apply (simp add: grewrite.intros(2) grewrite_append grewrite_ralts)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1108
  apply (simp add: hr_in_rstar hrewrite.intros(9))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1109
   apply (simp add: hr_in_rstar hrewrite.intros(10))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1110
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1111
  using hrewrite.intros(11) by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1112
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1113
lemma interleave_star1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1114
  shows "r h\<leadsto>* r' \<Longrightarrow> rder c r h\<leadsto>* rder c r'"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1115
  apply(induct rule : hrewrites.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1116
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1117
  by (meson hreal_trans interleave1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1118
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1119
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1120
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1121
lemma inside_simp_removal:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1122
  shows " rsimp (rder x (rsimp r)) = rsimp (rder x r)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1123
  apply(induct r)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1124
       apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1125
    apply(case_tac "rnullable r1")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1126
     apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1127
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1128
  using inside_simp_seq_nullable apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1129
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1130
  apply (smt (verit, del_insts) idiot2 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)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1131
   apply(subgoal_tac "rder x (RALTS xa) h\<leadsto>* rder x (rsimp (RALTS xa))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1132
  using hrewrites_simpeq apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1133
  using interleave_star1 simp_hrewrites apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1134
  by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1135
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1136
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1137
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1138
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1139
lemma rders_simp_same_simpders:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1140
  shows "s \<noteq> [] \<Longrightarrow> rders_simp r s = rsimp (rders r s)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1141
  apply(induct s rule: rev_induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1142
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1143
  apply(case_tac "xs = []")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1144
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1145
  apply(simp add: rders_append rders_simp_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1146
  using inside_simp_removal by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1147
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1148
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1149
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1150
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1151
lemma distinct_der:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1152
  shows "rsimp (rsimp_ALTs (map (rder x) (rdistinct rs {}))) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1153
         rsimp (rsimp_ALTs (rdistinct (map (rder x) rs) {}))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1154
  by (metis grewrites_simpalts gstar_rdistinct inside_simp_removal rder_rsimp_ALTs_commute)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1155
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1156
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1157
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1158
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1159
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1160
lemma rders_simp_lambda:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1161
  shows " rsimp \<circ> rder x \<circ> (\<lambda>r. rders_simp r xs) = (\<lambda>r. rders_simp r (xs @ [x]))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1162
  using rders_simp_append by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1163
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1164
lemma rders_simp_nonempty_simped:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1165
  shows "xs \<noteq> [] \<Longrightarrow> rsimp \<circ> (\<lambda>r. rders_simp r xs) = (\<lambda>r. rders_simp r xs)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1166
  using rders_simp_same_simpders rsimp_idem by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1167
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1168
lemma repeated_altssimp:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1169
  shows "\<forall>r \<in> set rs. rsimp r = r \<Longrightarrow> rsimp (rsimp_ALTs (rdistinct (rflts rs) {})) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1170
           rsimp_ALTs (rdistinct (rflts rs) {})"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1171
  by (metis map_idI rsimp.simps(2) rsimp_idem)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1172
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1173
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1174
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1175
lemma alts_closed_form: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1176
  shows "rsimp (rders_simp (RALTS rs) s) = rsimp (RALTS (map (\<lambda>r. rders_simp r s) rs))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1177
  apply(induct s rule: rev_induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1178
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1179
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1180
  apply(subst rders_simp_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1181
  apply(subgoal_tac " rsimp (rders_simp (rders_simp (RALTS rs) xs) [x]) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1182
 rsimp(rders_simp (rsimp_ALTs (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {})) [x])")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1183
   prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1184
  apply (metis inside_simp_removal rders_simp_one_char)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1185
  apply(simp only: )
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1186
  apply(subst rders_simp_one_char)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1187
  apply(subst rsimp_idem)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1188
  apply(subgoal_tac "rsimp (rder x (rsimp_ALTs (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {}))) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1189
                     rsimp ((rsimp_ALTs (map (rder x) (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {})))) ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1190
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1191
  using rder_rsimp_ALTs_commute apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1192
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1193
  apply(subgoal_tac "rsimp (rsimp_ALTs (map (rder x) (rdistinct (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)) {})))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1194
= rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {}))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1195
   prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1196
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1197
  using distinct_der apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1198
  apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1199
  apply(subgoal_tac " rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {})) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1200
                      rsimp (rsimp_ALTs (rdistinct ( (rflts (map (rder x) (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs)))) {}))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1201
   apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1202
  apply(subgoal_tac " rsimp (rsimp_ALTs (rdistinct (rflts (map (rder x) (map (rsimp \<circ> (\<lambda>r. rders_simp r xs)) rs))) {})) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1203
                      rsimp (rsimp_ALTs (rdistinct (rflts ( (map (rsimp \<circ> (rder x) \<circ> (\<lambda>r. rders_simp r xs)) rs))) {}))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1204
    apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1205
  apply(subst rders_simp_lambda)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1206
    apply(subst rders_simp_nonempty_simped)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1207
     apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1208
    apply(subgoal_tac "\<forall>r \<in> set  (map (\<lambda>r. rders_simp r (xs @ [x])) rs). rsimp r = r")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1209
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1210
     apply (simp add: rders_simp_same_simpders rsimp_idem)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1211
    apply(subst repeated_altssimp)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1212
     apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1213
  apply fastforce
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1214
   apply (metis inside_simp_removal list.map_comp rder.simps(4) rsimp.simps(2) rsimp_idem)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1215
  using simp_der_pierce_flts_prelim by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1216
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1217
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1218
lemma alts_closed_form_variant: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1219
  shows "s \<noteq> [] \<Longrightarrow> rders_simp (RALTS rs) s = rsimp (RALTS (map (\<lambda>r. rders_simp r s) rs))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1220
  by (metis alts_closed_form comp_apply rders_simp_nonempty_simped)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1221
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1222
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1223
lemma rsimp_seq_equal1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1224
  shows "rsimp_SEQ (rsimp r1) (rsimp r2) = rsimp_ALTs (rdistinct (rflts [rsimp_SEQ (rsimp r1) (rsimp r2)]) {})"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1225
  by (metis idem_after_simp1 rsimp.simps(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1226
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1227
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1228
fun sflat_aux :: "rrexp  \<Rightarrow> rrexp list " where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1229
  "sflat_aux (RALTS (r # rs)) = sflat_aux r @ rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1230
| "sflat_aux (RALTS []) = []"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1231
| "sflat_aux r = [r]"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1232
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1233
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1234
fun sflat :: "rrexp \<Rightarrow> rrexp" where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1235
  "sflat (RALTS (r # [])) = r"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1236
| "sflat (RALTS (r # rs)) = RALTS (sflat_aux r @ rs)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1237
| "sflat r = r"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1238
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1239
inductive created_by_seq:: "rrexp \<Rightarrow> bool" where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1240
  "created_by_seq (RSEQ r1 r2) "
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1241
| "created_by_seq r1 \<Longrightarrow> created_by_seq (RALT r1 r2)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1242
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1243
lemma seq_ders_shape1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1244
  shows "\<forall>r1 r2. \<exists>r3 r4. (rders (RSEQ r1 r2) s) = RSEQ r3 r4 \<or> (rders (RSEQ r1 r2) s) = RALT r3 r4"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1245
  apply(induct s rule: rev_induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1246
   apply auto[1]
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1247
  apply(rule allI)+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1248
  apply(subst rders_append)+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1249
  apply(subgoal_tac " \<exists>r3 r4. rders (RSEQ r1 r2) xs = RSEQ r3 r4 \<or> rders (RSEQ r1 r2) xs = RALT r3 r4 ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1250
   apply(erule exE)+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1251
   apply(erule disjE)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1252
    apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1253
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1254
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1255
lemma created_by_seq_der:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1256
  shows "created_by_seq r \<Longrightarrow> created_by_seq (rder c r)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1257
  apply(induct r)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1258
  apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1259
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1260
  using created_by_seq.cases apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1261
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1262
  apply (meson created_by_seq.cases rrexp.distinct(19) rrexp.distinct(21))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1263
  apply (metis created_by_seq.simps rder.simps(5))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1264
   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))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1265
  using created_by_seq.intros(1) by force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1266
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1267
lemma createdbyseq_left_creatable:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1268
  shows "created_by_seq (RALT r1 r2) \<Longrightarrow> created_by_seq r1"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1269
  using created_by_seq.cases by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1270
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1271
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1272
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1273
lemma recursively_derseq:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1274
  shows " created_by_seq (rders (RSEQ r1 r2) s)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1275
  apply(induct s rule: rev_induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1276
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1277
  using created_by_seq.intros(1) apply force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1278
  apply(subgoal_tac "created_by_seq (rders (RSEQ r1 r2) (xs @ [x]))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1279
  apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1280
  apply(subst rders_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1281
  apply(subgoal_tac "\<exists>r3 r4. rders (RSEQ r1 r2) xs = RSEQ r3 r4 \<or> 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1282
                    rders (RSEQ r1 r2) xs = RALT r3 r4")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1283
   prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1284
  using seq_ders_shape1 apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1285
  apply(erule exE)+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1286
  apply(erule disjE)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1287
   apply(subgoal_tac "created_by_seq (rders (RSEQ r3 r4) [x])")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1288
    apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1289
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1290
  using created_by_seq.intros(1) created_by_seq.intros(2) apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1291
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1292
  apply(subgoal_tac "created_by_seq r3")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1293
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1294
  using createdbyseq_left_creatable apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1295
  using created_by_seq.intros(2) created_by_seq_der by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1296
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1297
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1298
lemma recursively_derseq1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1299
  shows "r = rders (RSEQ r1 r2) s \<Longrightarrow> created_by_seq r"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1300
  using recursively_derseq by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1301
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1302
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1303
lemma sfau_head:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1304
  shows " created_by_seq r \<Longrightarrow> \<exists>ra rb rs. sflat_aux r = RSEQ ra rb # rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1305
  apply(induction r rule: created_by_seq.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1306
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1307
  by fastforce
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1308
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1309
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1310
lemma vsuf_prop1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1311
  shows "vsuf (xs @ [x]) r = (if (rnullable (rders r xs)) 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1312
                                then [x] # (map (\<lambda>s. s @ [x]) (vsuf xs r) )
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1313
                                else (map (\<lambda>s. s @ [x]) (vsuf xs r)) ) 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1314
             "
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1315
  apply(induct xs arbitrary: r)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1316
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1317
  apply(case_tac "rnullable r")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1318
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1319
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1320
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1321
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1322
fun  breakHead :: "rrexp list \<Rightarrow> rrexp list" where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1323
  "breakHead [] = [] "
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1324
| "breakHead (RALT r1 r2 # rs) = r1 # r2 # rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1325
| "breakHead (r # rs) = r # rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1326
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1327
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1328
lemma sfau_idem_der:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1329
  shows "created_by_seq r \<Longrightarrow> sflat_aux (rder c r) = breakHead (map (rder c) (sflat_aux r))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1330
  apply(induct rule: created_by_seq.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1331
   apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1332
  using sfau_head by fastforce
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1333
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1334
lemma vsuf_compose1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1335
  shows " \<not> rnullable (rders r1 xs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1336
       \<Longrightarrow> map (rder x \<circ> rders r2) (vsuf xs r1) = map (rders r2) (vsuf (xs @ [x]) r1)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1337
  apply(subst vsuf_prop1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1338
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1339
  by (simp add: rders_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1340
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1341
thm sfau_idem_der
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1342
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1343
lemma oneCharvsuf:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1344
  shows "breakHead [rder x (RSEQ r1 r2)] = RSEQ (rder x r1) r2 # map (rders r2)  (vsuf [x] r1)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1345
  by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1346
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1347
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1348
lemma vsuf_compose2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1349
  shows "(map (rders r2) (vsuf [x] (rders r1 xs))) @ map (rder x) (map (rders r2) (vsuf xs r1)) = map (rders r2) (vsuf (xs @ [x]) r1)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1350
proof(induct xs arbitrary: r1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1351
  case Nil
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1352
  then show ?case 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1353
    by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1354
next
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1355
  case (Cons a xs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1356
  have "rnullable (rders r1 xs) \<longrightarrow>        map (rders r2) (vsuf [x] (rders r1 (a # xs))) @ map (rder x) (map (rders r2) (vsuf (a # xs) r1)) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1357
       map (rders r2) (vsuf ((a # xs) @ [x]) r1)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1358
  proof
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1359
    assume nullableCond: "rnullable (rders r1 xs)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1360
    have "rnullable r1 \<longrightarrow> rder x (rders (rder a r2) xs) = rders (rder a r2) (xs @ [x])"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1361
        by (simp add: rders_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1362
    show " map (rders r2) (vsuf [x] (rders r1 (a # xs))) @ map (rder x) (map (rders r2) (vsuf (a # xs) r1)) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1363
       map (rders r2) (vsuf ((a # xs) @ [x]) r1)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1364
      using \<open>rnullable r1 \<longrightarrow> rder x (rders (rder a r2) xs) = rders (rder a r2) (xs @ [x])\<close> local.Cons by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1365
  qed
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1366
  then have "\<not> rnullable (rders r1 xs) \<longrightarrow> map (rders r2) (vsuf [x] (rders r1 (a # xs))) @ map (rder x) (map (rders r2) (vsuf (a # xs) r1)) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1367
            map (rders r2) (vsuf ((a # xs) @ [x]) r1)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1368
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1369
  by (smt (verit, ccfv_threshold) append_Cons append_Nil list.map_comp list.simps(8) list.simps(9) local.Cons rders.simps(1) rders.simps(2) rders_append vsuf.simps(1) vsuf.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1370
  then show ?case 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1371
    using \<open>rnullable (rders r1 xs) \<longrightarrow> map (rders r2) (vsuf [x] (rders r1 (a # xs))) @ map (rder x) (map (rders r2) (vsuf (a # xs) r1)) = map (rders r2) (vsuf ((a # xs) @ [x]) r1)\<close> by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1372
qed
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1373
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1374
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1375
lemma seq_sfau0:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1376
  shows  "sflat_aux (rders (RSEQ r1 r2) s) = (RSEQ (rders r1 s) r2) #
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1377
                                       (map (rders r2) (vsuf s r1)) "
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1378
proof(induct s rule: rev_induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1379
  case Nil
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1380
  then show ?case 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1381
    by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1382
next
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1383
  case (snoc x xs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1384
  then have LHS1:"sflat_aux (rders (RSEQ r1 r2) (xs @ [x]))  = sflat_aux (rder x (rders (RSEQ r1 r2) xs)) "
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1385
    by (simp add: rders_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1386
  then have LHS1A: "... = breakHead (map (rder x) (sflat_aux (rders (RSEQ r1 r2) xs)))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1387
    using recursively_derseq sfau_idem_der by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1388
  then have LHS1B: "... = breakHead (map (rder x)  (RSEQ (rders r1 xs) r2 # map (rders r2) (vsuf xs r1)))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1389
    using snoc by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1390
  then have LHS1C: "... = breakHead (rder x (RSEQ (rders r1 xs) r2) # map (rder x) (map (rders r2) (vsuf xs r1)))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1391
    by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1392
  then have LHS1D: "... = breakHead [rder x (RSEQ (rders r1 xs) r2)] @ map (rder x) (map (rders r2) (vsuf xs r1))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1393
    by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1394
  then have LHS1E: "... = RSEQ (rder x (rders r1 xs)) r2 # (map (rders r2) (vsuf [x] (rders r1 xs))) @ map (rder x) (map (rders r2) (vsuf xs r1))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1395
    by force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1396
  then have LHS1F: "... = RSEQ (rder x (rders r1 xs)) r2 # (map (rders r2) (vsuf (xs @ [x]) r1))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1397
    using vsuf_compose2 by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1398
  then have LHS1G: "... = RSEQ (rders r1 (xs @ [x])) r2 # (map (rders r2) (vsuf (xs @ [x]) r1))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1399
    using rders.simps(1) rders.simps(2) rders_append by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1400
  then show ?case
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1401
    using LHS1 LHS1A LHS1C LHS1D LHS1E LHS1F snoc by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1402
qed
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1403
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1404
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1405
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1406
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1407
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1408
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1409
lemma sflat_rsimpeq:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1410
  shows "created_by_seq r1 \<Longrightarrow> sflat_aux r1 =  rs \<Longrightarrow> rsimp r1 = rsimp (RALTS rs)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1411
  apply(induct r1 arbitrary: rs rule:  created_by_seq.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1412
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1413
  using rsimp_seq_equal1 apply force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1414
  by (metis head_one_more_simp rsimp.simps(2) sflat_aux.simps(1) simp_flatten)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1415
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1416
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1417
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1418
lemma seq_closed_form_general:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1419
  shows "rsimp (rders (RSEQ r1 r2) s) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1420
rsimp ( (RALTS ( (RSEQ (rders r1 s) r2 # (map (rders r2) (vsuf s r1))))))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1421
  apply(case_tac "s \<noteq> []")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1422
  apply(subgoal_tac "created_by_seq (rders (RSEQ r1 r2) s)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1423
  apply(subgoal_tac "sflat_aux (rders (RSEQ r1 r2) s) = RSEQ (rders r1 s) r2 # (map (rders r2) (vsuf s r1))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1424
  using sflat_rsimpeq apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1425
    apply (simp add: seq_sfau0)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1426
  using recursively_derseq1 apply blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1427
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1428
  by (metis idem_after_simp1 rsimp.simps(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1429
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1430
lemma seq_closed_form_aux1a:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1431
  shows "rsimp (RALTS (RSEQ (rders r1 s) r2 # rs)) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1432
           rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # rs))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1433
  by (metis head_one_more_simp rders.simps(1) rders_simp.simps(1) rders_simp_same_simpders rsimp.simps(1) rsimp_idem simp_flatten_aux0)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1434
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1435
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1436
lemma seq_closed_form_aux1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1437
  shows "rsimp (RALTS (RSEQ (rders r1 s) r2 # (map (rders r2) (vsuf s r1)))) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1438
           rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # (map (rders r2) (vsuf s r1))))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1439
  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)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1440
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1441
lemma add_simp_to_rest:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1442
  shows "rsimp (RALTS (r # rs)) = rsimp (RALTS (r # map rsimp rs))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1443
  by (metis append_Nil2 grewrite.intros(2) grewrite_simpalts head_one_more_simp list.simps(9) rsimp_ALTs.simps(2) spawn_simp_rsimpalts)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1444
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1445
lemma rsimp_compose_der2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1446
  shows "\<forall>s \<in> set ss. s \<noteq> [] \<Longrightarrow> map rsimp (map (rders r) ss) = map (\<lambda>s.  (rders_simp r s)) ss"  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1447
  by (simp add: rders_simp_same_simpders)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1448
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1449
lemma vsuf_nonempty:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1450
  shows "\<forall>s \<in> set ( vsuf s1 r). s \<noteq> []"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1451
  apply(induct s1 arbitrary: r)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1452
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1453
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1454
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1455
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1456
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1457
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1458
lemma seq_closed_form_aux2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1459
  shows "s \<noteq> [] \<Longrightarrow> rsimp ( (RALTS ( (RSEQ (rders_simp r1 s) r2 # (map (rders r2) (vsuf s r1)))))) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1460
         rsimp ( (RALTS ( (RSEQ (rders_simp r1 s) r2 # (map (rders_simp r2) (vsuf s r1))))))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1461
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1462
  by (metis add_simp_to_rest rsimp_compose_der2 vsuf_nonempty)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1463
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1464
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1465
lemma seq_closed_form: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1466
  shows "rsimp (rders_simp (RSEQ r1 r2) s) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1467
           rsimp (RALTS ((RSEQ (rders_simp r1 s) r2) # (map (rders_simp r2) (vsuf s r1))))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1468
proof (cases s)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1469
  case Nil
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1470
  then show ?thesis 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1471
    by (simp add: rsimp_seq_equal1[symmetric])
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1472
next
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1473
  case (Cons a list)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1474
  have "rsimp (rders_simp (RSEQ r1 r2) s) = rsimp (rsimp (rders (RSEQ r1 r2) s))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1475
    using local.Cons by (subst rders_simp_same_simpders)(simp_all)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1476
  also have "... = rsimp (rders (RSEQ r1 r2) s)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1477
    by (simp add: rsimp_idem)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1478
  also have "... = rsimp (RALTS (RSEQ (rders r1 s) r2 # map (rders r2) (vsuf s r1)))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1479
    using seq_closed_form_general by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1480
  also have "... = rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # map (rders r2) (vsuf s r1)))"  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1481
    by (simp only: seq_closed_form_aux1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1482
  also have "... = rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)))"  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1483
    using local.Cons by (subst seq_closed_form_aux2)(simp_all)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1484
  finally show ?thesis .
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1485
qed
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1486
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1487
lemma q: "s \<noteq> [] \<Longrightarrow> rders_simp (RSEQ r1 r2) s = rsimp (rders_simp (RSEQ r1 r2) s)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1488
  using rders_simp_same_simpders rsimp_idem by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1489
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1490
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1491
lemma seq_closed_form_variant: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1492
  assumes "s \<noteq> []"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1493
  shows "rders_simp (RSEQ r1 r2) s = 
543
b2bea5968b89 thesis_thys
Chengsong
parents: 513
diff changeset
  1494
             rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # 
b2bea5968b89 thesis_thys
Chengsong
parents: 513
diff changeset
  1495
        (map (rders_simp r2) (vsuf s r1))))"
505
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1496
  using assms q seq_closed_form by force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1497
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1498
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1499
fun hflat_aux :: "rrexp \<Rightarrow> rrexp list" where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1500
  "hflat_aux (RALT r1 r2) = hflat_aux r1 @ hflat_aux r2"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1501
| "hflat_aux r = [r]"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1502
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1503
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1504
fun hflat :: "rrexp \<Rightarrow> rrexp" where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1505
  "hflat (RALT r1 r2) = RALTS ((hflat_aux r1) @ (hflat_aux r2))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1506
| "hflat r = r"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1507
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1508
inductive created_by_star :: "rrexp \<Rightarrow> bool" where
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1509
  "created_by_star (RSEQ ra (RSTAR rb))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1510
| "\<lbrakk>created_by_star r1; created_by_star r2\<rbrakk> \<Longrightarrow> created_by_star (RALT r1 r2)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1511
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1512
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1513
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1514
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1515
lemma cbs_ders_cbs:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1516
  shows "created_by_star r \<Longrightarrow> created_by_star (rder c r)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1517
  apply(induct r rule: created_by_star.induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1518
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1519
  using created_by_star.intros(1) created_by_star.intros(2) apply auto[1]
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1520
  by (metis (mono_tags, lifting) created_by_star.simps list.simps(8) list.simps(9) rder.simps(4))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1521
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1522
lemma star_ders_cbs:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1523
  shows "created_by_star (rders (RSEQ r1 (RSTAR r2)) s)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1524
  apply(induct s rule: rev_induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1525
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1526
   apply (simp add: created_by_star.intros(1))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1527
  apply(subst rders_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1528
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1529
  using cbs_ders_cbs by auto
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1530
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1531
(*
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1532
lemma created_by_star_cases:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1533
  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 "
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1534
  by (meson created_by_star.cases)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1535
*)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1536
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1537
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1538
lemma hfau_pushin: 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1539
  shows "created_by_star r \<Longrightarrow> hflat_aux (rder c r) = concat (map hflat_aux (map (rder c) (hflat_aux r)))"
506
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1540
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1541
proof(induct r rule: created_by_star.induct)
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1542
  case (1 ra rb)
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1543
  then show ?case by simp
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1544
next
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1545
  case (2 r1 r2)
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1546
  then have "created_by_star (rder c r1)"
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1547
    using cbs_ders_cbs by blast
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1548
  then have "created_by_star (rder c r2)"
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1549
    using "2.hyps"(3) cbs_ders_cbs by auto
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1550
  then show ?case
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1551
    by (simp add: "2.hyps"(2) "2.hyps"(4))
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1552
  qed
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1553
513
ca7ca1f10f98 more isarfy
Chengsong
parents: 511
diff changeset
  1554
(*AALTS [a\x . b.c, b\x .c,  c \x]*)
ca7ca1f10f98 more isarfy
Chengsong
parents: 511
diff changeset
  1555
(*AALTS [a\x . b.c, AALTS [b\x .c, c\x]]*)
505
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1556
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1557
lemma stupdate_induct1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1558
  shows " concat (map (hflat_aux \<circ> (rder x \<circ> (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)))) Ss) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1559
          map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_update x r0 Ss)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1560
  apply(induct Ss)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1561
   apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1562
  by (simp add: rders_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1563
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1564
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1565
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1566
lemma stupdates_join_general:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1567
  shows  "concat (map hflat_aux (map (rder x) (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates xs r0 Ss)))) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1568
           map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates (xs @ [x]) r0 Ss)"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1569
  apply(induct xs arbitrary: Ss)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1570
   apply (simp)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1571
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1572
   apply auto[1]
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1573
  using stupdate_induct1 by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1574
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1575
lemma star_hfau_induct:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1576
  shows "hflat_aux (rders (RSEQ (rder c r0) (RSTAR r0)) s) =   
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1577
      map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates s r0 [[c]])"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1578
  apply(induct s rule: rev_induct)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1579
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1580
  apply(subst rders_append)+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1581
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1582
  apply(subst stupdates_append)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1583
  apply(subgoal_tac "created_by_star (rders (RSEQ (rder c r0) (RSTAR r0)) xs)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1584
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1585
  apply (simp add: star_ders_cbs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1586
  apply(subst hfau_pushin)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1587
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1588
  apply(subgoal_tac "concat (map hflat_aux (map (rder x) (hflat_aux (rders (RSEQ (rder c r0) (RSTAR r0)) xs)))) =
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1589
                     concat (map hflat_aux (map (rder x) ( map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates xs r0 [[c]])))) ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1590
   apply(simp only:)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1591
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1592
   apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1593
  apply(subst stupdates_append[symmetric])
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1594
  using stupdates_join_general by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1595
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1596
lemma starders_hfau_also1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1597
  shows "hflat_aux (rders (RSTAR r) (c # xs)) = map (\<lambda>s1. RSEQ (rders r s1) (RSTAR r)) (star_updates xs r [[c]])"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1598
  using star_hfau_induct by force
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1599
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1600
lemma hflat_aux_grewrites:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1601
  shows "a # rs \<leadsto>g* hflat_aux a @ rs"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1602
  apply(induct a arbitrary: rs)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1603
       apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1604
   apply(case_tac x)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1605
    apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1606
  apply(case_tac list)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1607
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1608
  apply (metis append.right_neutral append_Cons append_eq_append_conv2 grewrites.simps hflat_aux.simps(7) same_append_eq)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1609
   apply(case_tac lista)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1610
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1611
  apply (metis (no_types, lifting) append_Cons append_eq_append_conv2 gmany_steps_later greal_trans grewrite.intros(2) grewrites_append self_append_conv)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1612
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1613
  by simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1614
  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1615
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1616
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1617
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1618
lemma cbs_hfau_rsimpeq1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1619
  shows "rsimp (RALT a b) = rsimp (RALTS ((hflat_aux a) @ (hflat_aux b)))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1620
  apply(subgoal_tac "[a, b] \<leadsto>g* hflat_aux a @ hflat_aux b")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1621
  using grewrites_equal_rsimp apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1622
  by (metis append.right_neutral greal_trans grewrites_cons hflat_aux_grewrites)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1623
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1624
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1625
lemma hfau_rsimpeq2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1626
  shows "created_by_star r \<Longrightarrow> rsimp r = rsimp ( (RALTS (hflat_aux r)))"
506
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1627
  apply(induct rule: created_by_star.induct)
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1628
   apply simp
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1629
  apply (metis rsimp.simps(6) rsimp_seq_equal1)
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1630
  using cbs_hfau_rsimpeq1 hflat_aux.simps(1) by presburger
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1631
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1632
  
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1633
(*
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1634
lemma hfau_rsimpeq2_oldproof:  shows "created_by_star r \<Longrightarrow> rsimp r = rsimp ( (RALTS (hflat_aux r)))"
505
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1635
  apply(induct r)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1636
       apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1637
    apply (metis rsimp_seq_equal1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1638
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1639
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1640
  apply(case_tac x)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1641
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1642
  apply(case_tac "list")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1643
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1644
  apply (metis idem_after_simp1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1645
  apply(case_tac "lista")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1646
  prefer 2
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1647
   apply (metis hflat_aux.simps(8) idem_after_simp1 list.simps(8) list.simps(9) rsimp.simps(2))
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1648
  apply(subgoal_tac "rsimp (RALT a aa) = rsimp (RALTS (hflat_aux (RALT a aa)))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1649
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1650
  apply(subgoal_tac "rsimp (RALT a aa) = rsimp (RALTS (hflat_aux a @ hflat_aux aa))")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1651
  using hflat_aux.simps(1) apply presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1652
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1653
  using cbs_hfau_rsimpeq1 by fastforce
506
69ad05398894 thesis chapter 2 section 2.4 2.5
Chengsong
parents: 505
diff changeset
  1654
*)
505
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1655
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1656
lemma star_closed_form1:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1657
  shows "rsimp (rders (RSTAR r0) (c#s)) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1658
rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1659
  using hfau_rsimpeq2 rder.simps(6) rders.simps(2) star_ders_cbs starders_hfau_also1 by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1660
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1661
lemma star_closed_form2:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1662
  shows  "rsimp (rders_simp (RSTAR r0) (c#s)) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1663
rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1664
  by (metis list.distinct(1) rders_simp_same_simpders rsimp_idem star_closed_form1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1665
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1666
lemma star_closed_form3:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1667
  shows  "rsimp (rders_simp (RSTAR r0) (c#s)) =   (rders_simp (RSTAR r0) (c#s))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1668
  by (metis list.distinct(1) rders_simp_same_simpders star_closed_form1 star_closed_form2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1669
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1670
lemma star_closed_form4:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1671
  shows " (rders_simp (RSTAR r0) (c#s)) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1672
rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1673
  using star_closed_form2 star_closed_form3 by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1674
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1675
lemma star_closed_form5:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1676
  shows " rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rders r0 s1) (RSTAR r0) ) Ss         )))) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1677
          rsimp ( ( RALTS ( (map (\<lambda>s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss ))))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1678
  by (metis (mono_tags, lifting) list.map_comp map_eq_conv o_apply rsimp.simps(2) rsimp_idem)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1679
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1680
lemma star_closed_form6_hrewrites:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1681
  shows "  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1682
 (map (\<lambda>s1.  (RSEQ (rsimp (rders r0 s1)) (RSTAR r0)) ) Ss )
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1683
 scf\<leadsto>*
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1684
(map (\<lambda>s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss )"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1685
  apply(induct Ss)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1686
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1687
  apply (simp add: ss1)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1688
  by (metis (no_types, lifting) list.simps(9) rsimp.simps(1) rsimp_idem simp_hrewrites ss2)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1689
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1690
lemma star_closed_form6:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1691
  shows " rsimp ( ( RALTS ( (map (\<lambda>s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss )))) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1692
          rsimp ( ( RALTS ( (map (\<lambda>s1.  (RSEQ (rsimp (rders r0 s1)) (RSTAR r0)) ) Ss ))))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1693
  apply(subgoal_tac " map (\<lambda>s1.  (RSEQ (rsimp (rders r0 s1)) (RSTAR r0)) ) Ss  scf\<leadsto>*
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1694
                      map (\<lambda>s1.  rsimp (RSEQ  (rders r0 s1) (RSTAR r0)) ) Ss ")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1695
  using hrewrites_simpeq srewritescf_alt1 apply fastforce
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1696
  using star_closed_form6_hrewrites by blast
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1697
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1698
lemma stupdate_nonempty:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1699
  shows "\<forall>s \<in> set  Ss. s \<noteq> [] \<Longrightarrow> \<forall>s \<in> set (star_update c r Ss). s \<noteq> []"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1700
  apply(induct Ss)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1701
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1702
  apply(case_tac "rnullable (rders r a)")
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1703
   apply simp+
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1704
  done
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1705
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1706
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1707
lemma stupdates_nonempty:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1708
  shows "\<forall>s \<in> set Ss. s\<noteq> [] \<Longrightarrow> \<forall>s \<in> set (star_updates s r Ss). s \<noteq> []"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1709
  apply(induct s arbitrary: Ss)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1710
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1711
  apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1712
  using stupdate_nonempty by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1713
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1714
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1715
lemma star_closed_form8:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1716
  shows  
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1717
"rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ (rsimp (rders r0 s1)) (RSTAR r0) ) (star_updates s r0 [[c]]) ) ))) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1718
 rsimp ( ( RALTS ( (map (\<lambda>s1. RSEQ ( (rders_simp r0 s1)) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1719
  by (smt (verit, ccfv_SIG) list.simps(8) map_eq_conv rders__onechar rders_simp_same_simpders set_ConsD stupdates_nonempty)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1720
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1721
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1722
lemma star_closed_form:
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1723
  shows "rders_simp (RSTAR r0) (c#s) = 
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1724
rsimp ( RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) ))"
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1725
  apply(induct s)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1726
   apply simp
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1727
   apply (metis idem_after_simp1 rsimp.simps(1) rsimp.simps(6) rsimp_idem)
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1728
  using star_closed_form4 star_closed_form5 star_closed_form6 star_closed_form8 by presburger
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1729
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1730
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1731
unused_thms
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1732
5ce3bd8e5696 thesis chapter2 section 2.4
Chengsong
parents:
diff changeset
  1733
end