thys4/posix/BasicIdentities.thy
author Chengsong
Wed, 23 Aug 2023 03:02:31 +0100
changeset 668 3831621d7b14
parent 611 bc1df466150a
permissions -rw-r--r--
added technical Overview section, almost done introduction
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
587
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
     1
theory BasicIdentities 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
     2
  imports "Lexer" 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
     3
begin
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
     4
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
     5
datatype rrexp = 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
     6
  RZERO
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
     7
| RONE 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
     8
| RCHAR char
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
     9
| RSEQ rrexp rrexp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    10
| RALTS "rrexp list"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    11
| RSTAR rrexp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    12
| RNTIMES rrexp nat
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    13
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    14
abbreviation
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    15
  "RALT r1 r2 \<equiv> RALTS [r1, r2]"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    16
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    17
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    18
fun
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    19
 rnullable :: "rrexp \<Rightarrow> bool"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    20
where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    21
  "rnullable (RZERO) = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    22
| "rnullable (RONE) = True"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    23
| "rnullable (RCHAR c) = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    24
| "rnullable (RALTS rs) = (\<exists>r \<in> set rs. rnullable r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    25
| "rnullable (RSEQ r1 r2) = (rnullable r1 \<and> rnullable r2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    26
| "rnullable (RSTAR r) = True"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    27
| "rnullable (RNTIMES r n) = (if n = 0 then True else rnullable r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    28
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    29
fun
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    30
 rder :: "char \<Rightarrow> rrexp \<Rightarrow> rrexp"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    31
where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    32
  "rder c (RZERO) = RZERO"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    33
| "rder c (RONE) = RZERO"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    34
| "rder c (RCHAR d) = (if c = d then RONE else RZERO)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    35
| "rder c (RALTS rs) = RALTS (map (rder c) rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    36
| "rder c (RSEQ r1 r2) = 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    37
     (if rnullable r1
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    38
      then RALT   (RSEQ (rder c r1) r2) (rder c r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    39
      else RSEQ   (rder c r1) r2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    40
| "rder c (RSTAR r) = RSEQ  (rder c r) (RSTAR r)"   
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    41
| "rder c (RNTIMES r n) = (if n = 0 then RZERO else RSEQ (rder c r) (RNTIMES r (n - 1)))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    42
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    43
fun 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    44
  rders :: "rrexp \<Rightarrow> string \<Rightarrow> rrexp"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    45
where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    46
  "rders r [] = r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    47
| "rders r (c#s) = rders (rder c r) s"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    48
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    49
fun rdistinct :: "'a list \<Rightarrow>'a set \<Rightarrow> 'a list"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    50
  where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    51
  "rdistinct [] acc = []"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    52
| "rdistinct (x#xs)  acc = 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    53
     (if x \<in> acc then rdistinct xs  acc 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    54
      else x # (rdistinct xs  ({x} \<union> acc)))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    55
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    56
lemma rdistinct1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    57
  assumes "a \<in> acc"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    58
  shows "a \<notin> set (rdistinct rs acc)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    59
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    60
  apply(induct rs arbitrary: acc a)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    61
   apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    62
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    63
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    64
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    65
lemma rdistinct_does_the_job:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    66
  shows "distinct (rdistinct rs s)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    67
  apply(induct rs s rule: rdistinct.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    68
  apply(auto simp add: rdistinct1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    69
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    70
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    71
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    72
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    73
lemma rdistinct_concat:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    74
  assumes "set rs \<subseteq> rset"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    75
  shows "rdistinct (rs @ rsa) rset = rdistinct rsa rset"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    76
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    77
  apply(induct rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    78
  apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    79
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    80
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    81
lemma distinct_not_exist:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    82
  assumes "a \<notin> set rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    83
  shows "rdistinct rs rset = rdistinct rs (insert a rset)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    84
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    85
  apply(induct rs arbitrary: rset)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    86
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    87
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    88
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    89
lemma rdistinct_on_distinct:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    90
  shows "distinct rs \<Longrightarrow> rdistinct rs {} = rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    91
  apply(induct rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    92
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    93
  using distinct_not_exist by fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    94
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    95
lemma distinct_rdistinct_append:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    96
  assumes "distinct rs1" "\<forall>r \<in> set rs1. r \<notin> acc"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    97
  shows "rdistinct (rs1 @ rsa) acc = rs1 @ (rdistinct rsa (acc \<union> set rs1))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    98
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
    99
  apply(induct rs1 arbitrary: rsa acc)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   100
   apply(auto)[1]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   101
  apply(auto)[1]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   102
  apply(drule_tac x="rsa" in meta_spec)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   103
  apply(drule_tac x="{a} \<union> acc" in meta_spec)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   104
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   105
  apply(drule meta_mp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   106
   apply(auto)[1]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   107
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   108
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   109
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   110
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   111
lemma rdistinct_set_equality1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   112
  shows "set (rdistinct rs acc) = set rs - acc"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   113
  apply(induct rs acc rule: rdistinct.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   114
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   115
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   116
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   117
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   118
lemma rdistinct_set_equality:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   119
  shows "set (rdistinct rs {}) = set rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   120
  by (simp add: rdistinct_set_equality1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   121
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   122
611
Chengsong
parents: 587
diff changeset
   123
lemma distinct_removes_middle:
Chengsong
parents: 587
diff changeset
   124
  shows  "\<lbrakk>a \<in> set as\<rbrakk>
Chengsong
parents: 587
diff changeset
   125
    \<Longrightarrow> rdistinct (as @ as2) rset = rdistinct (as @ [a] @ as2) rset"
Chengsong
parents: 587
diff changeset
   126
and "rdistinct (ab # as @ [ab] @ as3) rset1 = rdistinct (ab # as @ as3) rset1"
Chengsong
parents: 587
diff changeset
   127
   apply(induct as arbitrary: rset rset1 ab as2 as3 a)
Chengsong
parents: 587
diff changeset
   128
     apply simp
Chengsong
parents: 587
diff changeset
   129
    apply simp
Chengsong
parents: 587
diff changeset
   130
   apply(case_tac "a \<in> rset")
Chengsong
parents: 587
diff changeset
   131
    apply simp
Chengsong
parents: 587
diff changeset
   132
    apply metis
Chengsong
parents: 587
diff changeset
   133
   apply simp
Chengsong
parents: 587
diff changeset
   134
   apply (metis insertI1)
Chengsong
parents: 587
diff changeset
   135
  apply(case_tac "a = ab")
Chengsong
parents: 587
diff changeset
   136
   apply simp
Chengsong
parents: 587
diff changeset
   137
   apply(case_tac "ab \<in> rset")
Chengsong
parents: 587
diff changeset
   138
    apply simp
Chengsong
parents: 587
diff changeset
   139
    apply presburger
Chengsong
parents: 587
diff changeset
   140
   apply (meson insertI1)
Chengsong
parents: 587
diff changeset
   141
  apply(case_tac "a \<in> rset")
Chengsong
parents: 587
diff changeset
   142
  apply (metis (no_types, opaque_lifting) Un_insert_left append_Cons insert_iff rdistinct.simps(2) sup_bot_left)
Chengsong
parents: 587
diff changeset
   143
  apply(case_tac "ab \<in> rset")
Chengsong
parents: 587
diff changeset
   144
  apply simp
Chengsong
parents: 587
diff changeset
   145
   apply (meson insert_iff)
Chengsong
parents: 587
diff changeset
   146
  apply simp
Chengsong
parents: 587
diff changeset
   147
  by (metis insertI1)
Chengsong
parents: 587
diff changeset
   148
Chengsong
parents: 587
diff changeset
   149
lemma distinct_removes_middle3:
Chengsong
parents: 587
diff changeset
   150
  shows  "\<lbrakk>a \<in> set as\<rbrakk>
Chengsong
parents: 587
diff changeset
   151
    \<Longrightarrow> rdistinct (as @ a #as2) rset = rdistinct (as @ as2) rset"
Chengsong
parents: 587
diff changeset
   152
  using distinct_removes_middle(1) by fastforce
Chengsong
parents: 587
diff changeset
   153
Chengsong
parents: 587
diff changeset
   154
lemma last_elem_out:
Chengsong
parents: 587
diff changeset
   155
  shows "\<lbrakk>x \<notin> set xs; x \<notin> rset \<rbrakk> \<Longrightarrow> rdistinct (xs @ [x]) rset = rdistinct xs rset @ [x]"
Chengsong
parents: 587
diff changeset
   156
  apply(induct xs arbitrary: rset)
Chengsong
parents: 587
diff changeset
   157
  apply simp+
Chengsong
parents: 587
diff changeset
   158
  done
Chengsong
parents: 587
diff changeset
   159
Chengsong
parents: 587
diff changeset
   160
Chengsong
parents: 587
diff changeset
   161
Chengsong
parents: 587
diff changeset
   162
Chengsong
parents: 587
diff changeset
   163
lemma rdistinct_concat_general:
Chengsong
parents: 587
diff changeset
   164
  shows "rdistinct (rs1 @ rs2) {} = (rdistinct rs1 {}) @ (rdistinct rs2 (set rs1))"
Chengsong
parents: 587
diff changeset
   165
  apply(induct rs1 arbitrary: rs2 rule: rev_induct)
Chengsong
parents: 587
diff changeset
   166
   apply simp
Chengsong
parents: 587
diff changeset
   167
  apply(drule_tac x = "x # rs2" in meta_spec)
Chengsong
parents: 587
diff changeset
   168
  apply simp
Chengsong
parents: 587
diff changeset
   169
  apply(case_tac "x \<in> set xs")
Chengsong
parents: 587
diff changeset
   170
   apply simp
Chengsong
parents: 587
diff changeset
   171
  
Chengsong
parents: 587
diff changeset
   172
   apply (simp add: distinct_removes_middle3 insert_absorb)
Chengsong
parents: 587
diff changeset
   173
  apply simp
Chengsong
parents: 587
diff changeset
   174
  by (simp add: last_elem_out)
Chengsong
parents: 587
diff changeset
   175
Chengsong
parents: 587
diff changeset
   176
Chengsong
parents: 587
diff changeset
   177
lemma distinct_once_enough:
Chengsong
parents: 587
diff changeset
   178
  shows "rdistinct (rs @ rsa) {} = rdistinct (rdistinct rs {} @ rsa) {}"
Chengsong
parents: 587
diff changeset
   179
  apply(subgoal_tac "distinct (rdistinct rs {})")
Chengsong
parents: 587
diff changeset
   180
   apply(subgoal_tac 
Chengsong
parents: 587
diff changeset
   181
" rdistinct (rdistinct rs {} @ rsa) {} = rdistinct rs {} @ (rdistinct rsa (set rs))")
Chengsong
parents: 587
diff changeset
   182
  apply(simp only:)
Chengsong
parents: 587
diff changeset
   183
  using rdistinct_concat_general apply blast
Chengsong
parents: 587
diff changeset
   184
  apply (simp add: distinct_rdistinct_append rdistinct_set_equality1)
Chengsong
parents: 587
diff changeset
   185
  by (simp add: rdistinct_does_the_job)
Chengsong
parents: 587
diff changeset
   186
  
Chengsong
parents: 587
diff changeset
   187
587
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   188
fun rflts :: "rrexp list \<Rightarrow> rrexp list"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   189
  where 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   190
  "rflts [] = []"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   191
| "rflts (RZERO # rs) = rflts rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   192
| "rflts ((RALTS rs1) # rs) = rs1 @ rflts rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   193
| "rflts (r1 # rs) = r1 # rflts rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   194
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   195
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   196
lemma rflts_def_idiot:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   197
  shows "\<lbrakk> a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1\<rbrakk> \<Longrightarrow> rflts (a # rs) = a # rflts rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   198
  apply(case_tac a)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   199
  apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   200
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   201
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   202
lemma rflts_def_idiot2:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   203
  shows "\<lbrakk>a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1; a \<in> set rs\<rbrakk> \<Longrightarrow> a \<in> set (rflts rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   204
  apply(induct rs rule: rflts.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   205
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   206
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   207
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   208
lemma flts_append:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   209
  shows "rflts (rs1 @ rs2) = rflts rs1 @ rflts rs2"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   210
  apply(induct rs1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   211
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   212
  apply(case_tac a)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   213
       apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   214
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   215
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   216
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   217
fun rsimp_ALTs :: " rrexp list \<Rightarrow> rrexp"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   218
  where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   219
  "rsimp_ALTs  [] = RZERO"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   220
| "rsimp_ALTs [r] =  r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   221
| "rsimp_ALTs rs = RALTS rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   222
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   223
lemma rsimpalts_conscons:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   224
  shows "rsimp_ALTs (r1 # rsa @ r2 # rsb) = RALTS (r1 # rsa @ r2 # rsb)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   225
  by (metis Nil_is_append_conv list.exhaust rsimp_ALTs.simps(3))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   226
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   227
lemma rsimp_alts_equal:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   228
  shows "rsimp_ALTs (rsa @ a # rsb @ a # rsc) = RALTS (rsa @ a # rsb @ a # rsc) "
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   229
  by (metis append_Cons append_Nil neq_Nil_conv rsimpalts_conscons)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   230
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   231
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   232
fun rsimp_SEQ :: " rrexp \<Rightarrow> rrexp \<Rightarrow> rrexp"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   233
  where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   234
  "rsimp_SEQ  RZERO _ = RZERO"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   235
| "rsimp_SEQ  _ RZERO = RZERO"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   236
| "rsimp_SEQ RONE r2 = r2"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   237
| "rsimp_SEQ r1 r2 = RSEQ r1 r2"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   238
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   239
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   240
fun rsimp :: "rrexp \<Rightarrow> rrexp" 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   241
  where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   242
  "rsimp (RSEQ r1 r2) = rsimp_SEQ  (rsimp r1) (rsimp r2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   243
| "rsimp (RALTS rs) = rsimp_ALTs  (rdistinct (rflts (map rsimp rs))  {}) "
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   244
| "rsimp r = r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   245
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   246
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   247
fun 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   248
  rders_simp :: "rrexp \<Rightarrow> string \<Rightarrow> rrexp"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   249
where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   250
  "rders_simp r [] = r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   251
| "rders_simp r (c#s) = rders_simp (rsimp (rder c r)) s"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   252
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   253
fun rsize :: "rrexp \<Rightarrow> nat" where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   254
  "rsize RZERO = 1"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   255
| "rsize (RONE) = 1" 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   256
| "rsize (RCHAR  c) = 1"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   257
| "rsize (RALTS  rs) = Suc (sum_list (map rsize rs))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   258
| "rsize (RSEQ  r1 r2) = Suc (rsize r1 + rsize r2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   259
| "rsize (RSTAR  r) = Suc (rsize r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   260
| "rsize (RNTIMES  r n) = Suc (rsize r) + n"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   261
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   262
abbreviation rsizes where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   263
  "rsizes rs \<equiv> sum_list (map rsize rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   264
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   265
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   266
lemma rder_rsimp_ALTs_commute:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   267
  shows "  (rder x (rsimp_ALTs rs)) = rsimp_ALTs (map (rder x) rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   268
  apply(induct rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   269
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   270
  apply(case_tac rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   271
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   272
  apply auto
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   273
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   274
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   275
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   276
lemma rsimp_aalts_smaller:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   277
  shows "rsize (rsimp_ALTs  rs) \<le> rsize (RALTS rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   278
  apply(induct rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   279
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   280
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   281
  apply(case_tac "rs = []")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   282
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   283
  apply(subgoal_tac "\<exists>rsp ap. rs = ap # rsp")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   284
   apply(erule exE)+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   285
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   286
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   287
  by(meson neq_Nil_conv)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   288
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   289
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   290
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   291
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   292
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   293
lemma rSEQ_mono:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   294
  shows "rsize (rsimp_SEQ r1 r2) \<le>rsize (RSEQ r1 r2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   295
  apply auto
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   296
  apply(induct r1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   297
       apply auto
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   298
      apply(case_tac "r2")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   299
       apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   300
     apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   301
          apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   302
     apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   303
         apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   304
     apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   305
        apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   306
     apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   307
         apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   308
  apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   309
         apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   310
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   311
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   312
lemma ralts_cap_mono:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   313
  shows "rsize (RALTS rs) \<le> Suc (rsizes rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   314
  by simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   315
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   316
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   317
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   318
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   319
lemma rflts_mono:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   320
  shows "rsizes (rflts rs) \<le> rsizes rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   321
  apply(induct rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   322
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   323
  apply(case_tac "a = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   324
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   325
  apply(case_tac "\<exists>rs1. a = RALTS rs1")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   326
  apply(erule exE)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   327
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   328
  apply(subgoal_tac "rflts (a # rs) = a # (rflts rs)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   329
   prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   330
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   331
  using rflts_def_idiot apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   332
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   333
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   334
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   335
lemma rdistinct_smaller: 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   336
  shows "rsizes (rdistinct rs ss) \<le> rsizes rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   337
  apply (induct rs arbitrary: ss)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   338
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   339
  by (simp add: trans_le_add2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   340
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   341
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   342
lemma rsimp_alts_mono :
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   343
  shows "\<And>x. (\<And>xa. xa \<in> set x \<Longrightarrow> rsize (rsimp xa) \<le> rsize xa)  \<Longrightarrow>
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   344
      rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {})) \<le> Suc (rsizes x)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   345
  apply(subgoal_tac "rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {} )) 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   346
                    \<le> rsize (RALTS (rdistinct (rflts (map rsimp x)) {} ))")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   347
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   348
  using rsimp_aalts_smaller apply auto[1]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   349
  apply(subgoal_tac "rsize (RALTS (rdistinct (rflts (map rsimp x)) {})) \<le>Suc (rsizes (rdistinct (rflts (map rsimp x)) {}))")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   350
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   351
  using ralts_cap_mono apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   352
  apply(subgoal_tac "rsizes (rdistinct (rflts (map rsimp x)) {}) \<le> rsizes (rflts (map rsimp x))")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   353
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   354
  using rdistinct_smaller apply presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   355
  apply(subgoal_tac "rsizes (rflts (map rsimp x)) \<le> rsizes (map rsimp x)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   356
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   357
  using rflts_mono apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   358
  apply(subgoal_tac "rsizes (map rsimp x) \<le> rsizes x")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   359
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   360
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   361
  apply (simp add: sum_list_mono)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   362
  by linarith
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   363
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   364
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   365
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   366
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   367
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   368
lemma rsimp_mono:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   369
  shows "rsize (rsimp r) \<le> rsize r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   370
  apply(induct r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   371
  apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   372
  apply(subgoal_tac "rsize (rsimp_SEQ (rsimp r1) (rsimp r2)) \<le> rsize (RSEQ (rsimp r1) (rsimp r2))")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   373
    apply force
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   374
  using rSEQ_mono
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   375
   apply presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   376
  using rsimp_alts_mono by auto
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   377
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   378
lemma idiot:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   379
  shows "rsimp_SEQ RONE r = r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   380
  apply(case_tac r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   381
       apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   382
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   383
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   384
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   385
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   386
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   387
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   388
lemma idiot2:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   389
  shows " \<lbrakk>r1 \<noteq> RZERO; r1 \<noteq> RONE;r2 \<noteq> RZERO\<rbrakk>
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   390
    \<Longrightarrow> rsimp_SEQ r1 r2 = RSEQ r1 r2"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   391
  apply(case_tac r1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   392
       apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   393
  apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   394
     apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   395
  apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   396
     apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   397
  apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   398
   apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   399
  apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   400
  apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   401
         apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   402
apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   403
         apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   404
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   405
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   406
lemma rders__onechar:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   407
  shows " (rders_simp r [c]) =  (rsimp (rders r [c]))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   408
  by simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   409
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   410
lemma rders_append:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   411
  "rders c (s1 @ s2) = rders (rders c s1) s2"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   412
  apply(induct s1 arbitrary: c s2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   413
  apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   414
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   415
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   416
lemma rders_simp_append:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   417
  "rders_simp c (s1 @ s2) = rders_simp (rders_simp c s1) s2"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   418
  apply(induct s1 arbitrary: c s2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   419
   apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   420
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   421
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   422
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   423
lemma rders_simp_one_char:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   424
  shows "rders_simp r [c] = rsimp (rder c r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   425
  apply auto
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   426
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   427
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   428
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   429
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   430
fun nonalt :: "rrexp \<Rightarrow> bool"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   431
  where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   432
  "nonalt (RALTS  rs) = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   433
| "nonalt r = True"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   434
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   435
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   436
fun good :: "rrexp \<Rightarrow> bool" where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   437
  "good RZERO = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   438
| "good (RONE) = True" 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   439
| "good (RCHAR c) = True"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   440
| "good (RALTS []) = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   441
| "good (RALTS [r]) = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   442
| "good (RALTS (r1 # r2 # rs)) = ((distinct ( (r1 # r2 # rs))) \<and>(\<forall>r' \<in> set (r1 # r2 # rs). good r' \<and> nonalt r'))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   443
| "good (RSEQ RZERO _) = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   444
| "good (RSEQ RONE _) = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   445
| "good (RSEQ  _ RZERO) = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   446
| "good (RSEQ r1 r2) = (good r1 \<and> good r2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   447
| "good (RSTAR r) = True"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   448
| "good (RNTIMES r n) = True"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   449
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   450
lemma  k0a:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   451
  shows "rflts [RALTS rs] =   rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   452
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   453
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   454
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   455
lemma bbbbs:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   456
  assumes "good r" "r = RALTS rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   457
  shows "rsimp_ALTs  (rflts [r]) = RALTS rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   458
  using  assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   459
  by (metis good.simps(4) good.simps(5) k0a rsimp_ALTs.elims)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   460
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   461
lemma bbbbs1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   462
  shows "nonalt r \<or> (\<exists> rs. r  = RALTS  rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   463
  by (meson nonalt.elims(3))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   464
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   465
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   466
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   467
lemma good0:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   468
  assumes "rs \<noteq> Nil" "\<forall>r \<in> set rs. nonalt r" "distinct rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   469
  shows "good (rsimp_ALTs rs) \<longleftrightarrow> (\<forall>r \<in> set rs. good r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   470
  using  assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   471
  apply(induct  rs rule: rsimp_ALTs.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   472
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   473
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   474
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   475
lemma flts1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   476
  assumes "good r" 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   477
  shows "rflts [r] \<noteq> []"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   478
  using  assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   479
  apply(induct r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   480
       apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   481
  using good.simps(4) by blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   482
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   483
lemma flts2:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   484
  assumes "good r" 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   485
  shows "\<forall>r' \<in> set (rflts [r]). good r' \<and> nonalt r'"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   486
  using  assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   487
  apply(induct r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   488
       apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   489
      apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   490
     apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   491
    prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   492
    apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   493
    apply(auto)[1]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   494
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   495
     apply (metis flts1 good.simps(5) good.simps(6) k0a neq_Nil_conv)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   496
    apply (metis flts1 good.simps(5) good.simps(6) k0a neq_Nil_conv)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   497
   apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   498
   apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   499
  by simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   500
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   501
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   502
lemma flts3:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   503
  assumes "\<forall>r \<in> set rs. good r \<or> r = RZERO" 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   504
  shows "\<forall>r \<in> set (rflts rs). good r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   505
  using  assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   506
  apply(induct rs arbitrary: rule: rflts.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   507
        apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   508
  by (metis UnE flts2 k0a)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   509
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   510
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   511
lemma  k0:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   512
  shows "rflts (r # rs1) = rflts [r] @ rflts rs1"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   513
  apply(induct r arbitrary: rs1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   514
   apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   515
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   516
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   517
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   518
lemma good_SEQ:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   519
  assumes "r1 \<noteq> RZERO" "r2 \<noteq> RZERO" " r1 \<noteq> RONE"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   520
  shows "good (RSEQ r1 r2) = (good r1 \<and> good r2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   521
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   522
  apply(case_tac r1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   523
       apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   524
  apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   525
          apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   526
  apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   527
         apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   528
  apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   529
        apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   530
  apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   531
         apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   532
apply(case_tac r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   533
         apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   534
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   535
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   536
lemma rsize0:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   537
  shows "0 < rsize r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   538
  apply(induct  r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   539
       apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   540
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   541
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   542
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   543
fun nonnested :: "rrexp \<Rightarrow> bool"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   544
  where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   545
  "nonnested (RALTS []) = True"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   546
| "nonnested (RALTS ((RALTS rs1) # rs2)) = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   547
| "nonnested (RALTS (r # rs2)) = nonnested (RALTS rs2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   548
| "nonnested r = True"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   549
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   550
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   551
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   552
lemma  k00:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   553
  shows "rflts (rs1 @ rs2) = rflts rs1 @ rflts rs2"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   554
  apply(induct rs1 arbitrary: rs2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   555
   apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   556
  by (metis append.assoc k0)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   557
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   558
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   559
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   560
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   561
lemma k0b:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   562
  assumes "nonalt r" "r \<noteq> RZERO"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   563
  shows "rflts [r] = [r]"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   564
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   565
  apply(case_tac  r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   566
  apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   567
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   568
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   569
lemma nn1qq:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   570
  assumes "nonnested (RALTS rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   571
  shows "\<nexists> rs1. RALTS rs1 \<in> set rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   572
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   573
  apply(induct rs rule: rflts.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   574
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   575
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   576
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   577
 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   578
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   579
lemma n0:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   580
  shows "nonnested (RALTS rs) \<longleftrightarrow> (\<forall>r \<in> set rs. nonalt r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   581
  apply(induct rs )
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   582
   apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   583
    apply (metis list.set_intros(1) nn1qq nonalt.elims(3))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   584
  apply (metis nonalt.elims(2) nonnested.simps(3) nonnested.simps(4) nonnested.simps(5) nonnested.simps(6) nonnested.simps(7) nonnested.simps(8))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   585
  using bbbbs1 apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   586
  by (metis bbbbs1 list.set_intros(2) nn1qq)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   587
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   588
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   589
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   590
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   591
lemma nn1c:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   592
  assumes "\<forall>r \<in> set rs. nonnested r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   593
  shows "\<forall>r \<in> set (rflts rs). nonalt r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   594
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   595
  apply(induct rs rule: rflts.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   596
        apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   597
  using n0 by blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   598
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   599
lemma nn1bb:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   600
  assumes "\<forall>r \<in> set rs. nonalt r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   601
  shows "nonnested (rsimp_ALTs  rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   602
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   603
  apply(induct  rs rule: rsimp_ALTs.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   604
    apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   605
  using nonalt.simps(1) nonnested.elims(3) apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   606
  using n0 by auto
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   607
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   608
lemma bsimp_ASEQ0:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   609
  shows "rsimp_SEQ  r1 RZERO = RZERO"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   610
  apply(induct r1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   611
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   612
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   613
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   614
lemma nn1b:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   615
  shows "nonnested (rsimp r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   616
  apply(induct r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   617
       apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   618
  apply(case_tac "rsimp r1 = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   619
    apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   620
 apply(case_tac "rsimp r2 = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   621
   apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   622
    apply(subst bsimp_ASEQ0)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   623
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   624
  apply(case_tac "\<exists>bs. rsimp r1 = RONE")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   625
    apply(auto)[1]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   626
  using idiot apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   627
  apply (simp add: idiot2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   628
  by (metis (mono_tags, lifting) image_iff list.set_map nn1bb nn1c rdistinct_set_equality)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   629
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   630
lemma nonalt_flts_rd:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   631
  shows "\<lbrakk>xa \<in> set (rdistinct (rflts (map rsimp rs)) {})\<rbrakk>
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   632
       \<Longrightarrow> nonalt xa"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   633
  by (metis Diff_empty ex_map_conv nn1b nn1c rdistinct_set_equality1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   634
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   635
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   636
lemma rsimpalts_implies1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   637
  shows " rsimp_ALTs (a # rdistinct rs {a}) = RZERO \<Longrightarrow> a = RZERO"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   638
  using rsimp_ALTs.elims by auto
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   639
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   640
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   641
lemma rsimpalts_implies2:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   642
  shows "rsimp_ALTs (a # rdistinct rs rset) = RZERO \<Longrightarrow> rdistinct rs rset = []"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   643
  by (metis append_butlast_last_id rrexp.distinct(7) rsimpalts_conscons)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   644
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   645
lemma rsimpalts_implies21:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   646
  shows "rsimp_ALTs (a # rdistinct rs {a}) = RZERO \<Longrightarrow> rdistinct rs {a} = []"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   647
  using rsimpalts_implies2 by blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   648
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   649
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   650
lemma bsimp_ASEQ2:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   651
  shows "rsimp_SEQ RONE r2 =  r2"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   652
  apply(induct r2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   653
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   654
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   655
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   656
lemma elem_smaller_than_set:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   657
  shows "xa \<in> set  list \<Longrightarrow> rsize xa < Suc (rsizes list)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   658
  apply(induct list)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   659
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   660
  by (metis image_eqI le_imp_less_Suc list.set_map member_le_sum_list)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   661
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   662
lemma rsimp_list_mono:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   663
  shows "rsizes (map rsimp rs) \<le> rsizes rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   664
  apply(induct rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   665
   apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   666
  by (simp add: add_mono_thms_linordered_semiring(1) rsimp_mono)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   667
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   668
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   669
(*says anything coming out of simp+flts+db will be good*)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   670
lemma good2_obv_simplified:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   671
  shows " \<lbrakk>\<forall>y. rsize y < Suc (rsizes rs) \<longrightarrow> good (rsimp y) \<or> rsimp y = RZERO;
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   672
           xa \<in> set (rdistinct (rflts (map rsimp rs)) {}); good (rsimp xa) \<or> rsimp xa = RZERO\<rbrakk> \<Longrightarrow> good xa"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   673
  apply(subgoal_tac " \<forall>xa' \<in> set (map rsimp rs). good xa' \<or> xa' = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   674
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   675
   apply (simp add: elem_smaller_than_set)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   676
  by (metis Diff_empty flts3 rdistinct_set_equality1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   677
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   678
thm Diff_empty flts3 rdistinct_set_equality1
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   679
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   680
lemma good1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   681
  shows "good (rsimp a) \<or> rsimp a = RZERO"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   682
  apply(induct a taking: rsize rule: measure_induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   683
  apply(case_tac x)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   684
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   685
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   686
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   687
  prefer 3
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   688
    apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   689
   prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   690
   apply(simp only:)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   691
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   692
  apply (smt (verit, ccfv_threshold) add_mono_thms_linordered_semiring(1) elem_smaller_than_set good0 good2_obv_simplified le_eq_less_or_eq nonalt_flts_rd order_less_trans plus_1_eq_Suc rdistinct_does_the_job rdistinct_smaller rflts_mono rsimp_ALTs.simps(1) rsimp_list_mono)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   693
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   694
  apply(subgoal_tac "good (rsimp x41) \<or> rsimp x41 = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   695
   apply(subgoal_tac "good (rsimp x42) \<or> rsimp x42 = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   696
    apply(case_tac "rsimp x41 = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   697
     apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   698
    apply(case_tac "rsimp x42 = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   699
     apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   700
  using bsimp_ASEQ0 apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   701
    apply(subgoal_tac "good (rsimp x41)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   702
     apply(subgoal_tac "good (rsimp x42)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   703
      apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   704
  apply (metis bsimp_ASEQ2 good_SEQ idiot2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   705
  apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   706
  apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   707
  using less_add_Suc2 apply blast  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   708
  using less_iff_Suc_add apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   709
  using good.simps(45) rsimp.simps(7) by presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   710
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   711
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   712
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   713
fun
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   714
  RL :: "rrexp \<Rightarrow> string set"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   715
where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   716
  "RL (RZERO) = {}"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   717
| "RL (RONE) = {[]}"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   718
| "RL (RCHAR c) = {[c]}"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   719
| "RL (RSEQ r1 r2) = (RL r1) ;; (RL r2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   720
| "RL (RALTS rs) = (\<Union> (set (map RL rs)))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   721
| "RL (RSTAR r) = (RL r)\<star>"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   722
| "RL (RNTIMES r n) = (RL r) ^^ n"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   723
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   724
lemma pow_rempty_iff:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   725
  shows "[] \<in> (RL r) ^^ n \<longleftrightarrow> (if n = 0 then True else [] \<in> (RL r))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   726
  by (induct n) (auto simp add: Sequ_def)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   727
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   728
lemma RL_rnullable:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   729
  shows "rnullable r = ([] \<in> RL r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   730
  apply(induct r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   731
        apply(auto simp add: Sequ_def pow_rempty_iff)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   732
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   733
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   734
lemma concI_if_Nil1: "[] \<in> A \<Longrightarrow> xs : B \<Longrightarrow> xs \<in> A ;; B"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   735
by (metis append_Nil concI)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   736
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   737
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   738
lemma empty_pow_add:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   739
  fixes A::"string set"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   740
  assumes "[] \<in> A" "s \<in> A ^^ n"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   741
  shows "s \<in> A ^^ (n + m)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   742
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   743
  apply(induct m arbitrary: n)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   744
   apply(auto simp add: Sequ_def)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   745
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   746
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   747
(*
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   748
lemma der_pow:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   749
  shows "Der c (A ^^ n) = (if n = 0 then {} else (Der c A) ;; (A ^^ (n - 1)))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   750
  apply(induct n arbitrary: A)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   751
   apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   752
  by (smt (verit, best) Suc_pred concE concI concI_if_Nil2 conc_pow_comm lang_pow.simps(2))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   753
*)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   754
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   755
lemma RL_rder:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   756
  shows "RL (rder c r) = Der c (RL r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   757
  apply(induct r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   758
  apply(auto simp add: Sequ_def Der_def)[5]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   759
        apply (metis append_Cons)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   760
  using RL_rnullable apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   761
  apply (metis append_eq_Cons_conv)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   762
  apply (metis append_Cons)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   763
    apply (metis RL_rnullable append_eq_Cons_conv)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   764
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   765
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   766
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   767
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   768
lemma RL_rsimp_RSEQ:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   769
  shows "RL (rsimp_SEQ r1 r2) = (RL r1 ;; RL r2)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   770
  apply(induct r1 r2 rule: rsimp_SEQ.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   771
  apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   772
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   773
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   774
lemma RL_rsimp_RALTS:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   775
  shows "RL (rsimp_ALTs rs) = (\<Union> (set (map RL rs)))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   776
  apply(induct rs rule: rsimp_ALTs.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   777
  apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   778
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   779
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   780
lemma RL_rsimp_rdistinct:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   781
  shows "(\<Union> (set (map RL (rdistinct rs {})))) = (\<Union> (set (map RL rs)))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   782
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   783
  apply (metis Diff_iff rdistinct_set_equality1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   784
  by (metis Diff_empty rdistinct_set_equality1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   785
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   786
lemma RL_rsimp_rflts:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   787
  shows "(\<Union> (set (map RL (rflts rs)))) = (\<Union> (set (map RL rs)))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   788
  apply(induct rs rule: rflts.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   789
  apply(simp_all)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   790
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   791
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   792
lemma RL_rsimp:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   793
  shows "RL r = RL (rsimp r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   794
  apply(induct r rule: rsimp.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   795
       apply(auto simp add: Sequ_def RL_rsimp_RSEQ)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   796
  using RL_rsimp_RALTS RL_rsimp_rdistinct RL_rsimp_rflts apply auto[1]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   797
  by (smt (verit, del_insts) RL_rsimp_RALTS RL_rsimp_rdistinct RL_rsimp_rflts UN_E image_iff list.set_map)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   798
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   799
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   800
lemma qqq1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   801
  shows "RZERO \<notin> set (rflts (map rsimp rs))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   802
  by (metis ex_map_conv flts3 good.simps(1) good1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   803
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   804
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   805
fun nonazero :: "rrexp \<Rightarrow> bool"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   806
  where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   807
  "nonazero RZERO = False"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   808
| "nonazero r = True"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   809
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   810
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   811
lemma flts_single1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   812
  assumes "nonalt r" "nonazero r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   813
  shows "rflts [r] = [r]"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   814
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   815
  apply(induct r)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   816
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   817
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   818
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   819
lemma nonalt0_flts_keeps:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   820
  shows "(a \<noteq> RZERO) \<and> (\<forall>rs. a \<noteq> RALTS rs) \<Longrightarrow> rflts (a # xs) = a # rflts xs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   821
  apply(case_tac a)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   822
       apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   823
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   824
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   825
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   826
lemma nonalt0_fltseq:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   827
  shows "\<forall>r \<in> set rs. r \<noteq> RZERO \<and> nonalt r \<Longrightarrow> rflts rs = rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   828
  apply(induct rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   829
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   830
  apply(case_tac "a = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   831
   apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   832
  apply(case_tac "\<exists>rs1. a = RALTS rs1")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   833
   apply(erule exE)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   834
   apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   835
  using nonalt0_flts_keeps by presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   836
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   837
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   838
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   839
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   840
lemma goodalts_nonalt:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   841
  shows "good (RALTS rs) \<Longrightarrow> rflts rs = rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   842
  apply(induct x == "RALTS rs" arbitrary: rs rule: good.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   843
    apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   844
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   845
  using good.simps(5) apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   846
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   847
  apply(case_tac "r1 = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   848
  using good.simps(1) apply force
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   849
  apply(case_tac "r2 = RZERO")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   850
  using good.simps(1) apply force
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   851
  apply(subgoal_tac "rflts (r1 # r2 # rs) = r1 # r2 # rflts rs")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   852
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   853
   apply (metis nonalt.simps(1) rflts_def_idiot)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   854
  apply(subgoal_tac "\<forall>r \<in> set rs. r \<noteq> RZERO \<and> nonalt r")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   855
   apply(subgoal_tac "rflts rs = rs")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   856
    apply presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   857
  using nonalt0_fltseq apply presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   858
  using good.simps(1) by blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   859
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   860
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   861
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   862
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   863
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   864
lemma test:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   865
  assumes "good r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   866
  shows "rsimp r = r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   867
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   868
  using assms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   869
  apply(induct rule: good.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   870
                      apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   871
                      apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   872
                      apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   873
                      apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   874
                      apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   875
                      apply(subgoal_tac "distinct (r1 # r2 # rs)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   876
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   877
  using good.simps(6) apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   878
  apply(subgoal_tac "rflts (r1 # r2 # rs ) = r1 # r2 # rs")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   879
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   880
  using goodalts_nonalt apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   881
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   882
                      apply(subgoal_tac "r1 \<noteq> r2")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   883
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   884
                      apply (meson distinct_length_2_or_more)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   885
                      apply(subgoal_tac "r1 \<notin> set rs")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   886
                      apply(subgoal_tac "r2 \<notin> set rs")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   887
                      apply(subgoal_tac "\<forall>r \<in> set rs. rsimp r = r")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   888
                      apply(subgoal_tac "map rsimp rs = rs")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   889
  apply simp             
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   890
                      apply(subgoal_tac "\<forall>r \<in>  {r1, r2}. r \<notin> set rs")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   891
  apply (metis distinct_not_exist rdistinct_on_distinct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   892
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   893
                      apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   894
                      apply (meson map_idI)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   895
                      apply (metis good.simps(6) insert_iff list.simps(15))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   896
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   897
  apply (meson distinct.simps(2))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   898
                      apply (simp add: distinct_length_2_or_more)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   899
                      apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   900
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   901
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   902
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   903
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   904
lemma rsimp_idem:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   905
  shows "rsimp (rsimp r) = rsimp r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   906
  using test good1
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   907
  by force
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   908
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   909
corollary rsimp_inner_idem4:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   910
  shows "rsimp r = RALTS rs \<Longrightarrow> rflts rs = rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   911
  by (metis good1 goodalts_nonalt rrexp.simps(12))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   912
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   913
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   914
lemma head_one_more_simp:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   915
  shows "map rsimp (r # rs) = map rsimp (( rsimp r) # rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   916
  by (simp add: rsimp_idem)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   917
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   918
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   919
lemma der_simp_nullability:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   920
  shows "rnullable r = rnullable (rsimp r)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   921
  using RL_rnullable RL_rsimp by auto
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   922
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   923
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   924
lemma no_alt_short_list_after_simp:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   925
  shows "RALTS rs = rsimp r \<Longrightarrow> rsimp_ALTs rs = RALTS rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   926
  by (metis bbbbs good1 k0a rrexp.simps(12))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   927
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   928
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   929
lemma no_further_dB_after_simp:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   930
  shows "RALTS rs = rsimp r \<Longrightarrow> rdistinct rs {} = rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   931
  apply(subgoal_tac "good (RALTS rs)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   932
  apply(subgoal_tac "distinct rs")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   933
  using rdistinct_on_distinct apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   934
  apply (metis distinct.simps(1) distinct.simps(2) empty_iff good.simps(6) list.exhaust set_empty2)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   935
  using good1 by fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   936
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   937
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   938
lemma idem_after_simp1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   939
  shows "rsimp_ALTs (rdistinct (rflts [rsimp aa]) {}) = rsimp aa"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   940
  apply(case_tac "rsimp aa")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   941
  apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   942
  apply (metis no_alt_short_list_after_simp no_further_dB_after_simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   943
   apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   944
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   945
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   946
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   947
lemma identity_wwo0:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   948
  shows "rsimp (rsimp_ALTs (RZERO # rs)) = rsimp (rsimp_ALTs rs)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   949
  apply (metis idem_after_simp1 list.exhaust list.simps(8) list.simps(9) rflts.simps(2) rsimp.simps(2) rsimp.simps(3) rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   950
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   951
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   952
lemma distinct_removes_last:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   953
  shows "\<lbrakk>a \<in> set as\<rbrakk>
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   954
    \<Longrightarrow> rdistinct as rset = rdistinct (as @ [a]) rset"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   955
and "rdistinct (ab # as @ [ab]) rset1 = rdistinct (ab # as) rset1"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   956
  apply(induct as arbitrary: rset ab rset1 a)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   957
     apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   958
    apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   959
  apply(case_tac "aa \<in> rset")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   960
   apply(case_tac "a = aa")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   961
  apply (metis append_Cons)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   962
    apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   963
   apply(case_tac "a \<in> set as")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   964
  apply (metis append_Cons rdistinct.simps(2) set_ConsD)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   965
   apply(case_tac "a = aa")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   966
    prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   967
    apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   968
   apply (metis append_Cons)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   969
  apply(case_tac "ab \<in> rset1")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   970
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   971
   apply(subgoal_tac "rdistinct (ab # (a # as) @ [ab]) rset1 = 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   972
               ab # (rdistinct ((a # as) @ [ab]) (insert ab rset1))")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   973
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   974
  apply force
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   975
  apply(simp only:)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   976
     apply(subgoal_tac "rdistinct (ab # a # as) rset1 = ab # (rdistinct (a # as) (insert ab rset1))")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   977
    apply(simp only:)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   978
    apply(subgoal_tac "rdistinct ((a # as) @ [ab]) (insert ab rset1) = rdistinct (a # as) (insert ab rset1)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   979
     apply blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   980
    apply(case_tac "a \<in> insert ab rset1")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   981
     apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   982
     apply (metis insertI1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   983
    apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   984
    apply (meson insertI1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   985
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   986
  apply(subgoal_tac "rdistinct ((a # as) @ [ab]) rset1 = rdistinct (a # as) rset1")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   987
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   988
  by (metis append_Cons insert_iff insert_is_Un rdistinct.simps(2))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   989
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   990
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   991
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   992
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   993
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   994
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   995
lemma distinct_removes_list:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   996
  shows "\<lbrakk> \<forall>r \<in> set rs. r \<in> set as\<rbrakk> \<Longrightarrow> rdistinct (as @ rs) {} = rdistinct as {}"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   997
  apply(induct rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   998
   apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
   999
  apply(subgoal_tac "rdistinct (as @ a # rs) {} = rdistinct (as @ rs) {}")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1000
   prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1001
  apply (metis append_Cons append_Nil distinct_removes_middle(1))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1002
  by presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1003
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1004
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1005
lemma spawn_simp_rsimpalts:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1006
  shows "rsimp (rsimp_ALTs rs) = rsimp (rsimp_ALTs (map rsimp rs))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1007
  apply(cases rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1008
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1009
  apply(case_tac list)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1010
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1011
   apply(subst rsimp_idem[symmetric])
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1012
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1013
  apply(subgoal_tac "rsimp_ALTs rs = RALTS rs")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1014
   apply(simp only:)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1015
   apply(subgoal_tac "rsimp_ALTs (map rsimp rs) = RALTS (map rsimp rs)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1016
    apply(simp only:)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1017
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1018
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1019
   prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1020
  using rsimp_ALTs.simps(3) apply presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1021
  apply auto
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1022
  apply(subst rsimp_idem)+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1023
  by (metis comp_apply rsimp_idem)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1024
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1025
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1026
lemma simp_singlealt_flatten:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1027
  shows "rsimp (RALTS [RALTS rsa]) = rsimp (RALTS (rsa @ []))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1028
  apply(induct rsa)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1029
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1030
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1031
  by (metis idem_after_simp1 list.simps(9) rsimp.simps(2))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1032
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1033
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1034
lemma good1_rsimpalts:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1035
  shows "rsimp r = RALTS rs \<Longrightarrow> rsimp_ALTs rs = RALTS rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1036
  by (metis no_alt_short_list_after_simp) 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1037
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1038
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1039
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1040
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1041
lemma good1_flatten:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1042
  shows "\<lbrakk> rsimp r =  (RALTS rs1)\<rbrakk>
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1043
       \<Longrightarrow> rflts (rsimp_ALTs rs1 # map rsimp rsb) = rflts (rs1 @ map rsimp rsb)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1044
  apply(subst good1_rsimpalts)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1045
   apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1046
  apply(subgoal_tac "rflts (rs1 @ map rsimp rsb) = rs1 @ rflts (map rsimp rsb)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1047
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1048
  using flts_append rsimp_inner_idem4 by presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1049
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1050
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1051
lemma flatten_rsimpalts:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1052
  shows "rflts (rsimp_ALTs (rdistinct (rflts (map rsimp rsa)) {}) # map rsimp rsb) = 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1053
         rflts ( (rdistinct (rflts (map rsimp rsa)) {}) @ map rsimp rsb)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1054
  apply(case_tac "map rsimp rsa")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1055
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1056
  apply(case_tac "list")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1057
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1058
   apply(case_tac a)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1059
        apply simp+
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1060
    apply(rename_tac rs1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1061
    apply (metis good1_flatten map_eq_Cons_D no_further_dB_after_simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1062
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1063
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1064
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1065
  apply(subgoal_tac "\<forall>r \<in> set( rflts (map rsimp rsa)). good r")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1066
   apply(case_tac "rdistinct (rflts (map rsimp rsa)) {}")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1067
     apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1068
  apply auto[1]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1069
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1070
  apply(simp)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1071
   apply(case_tac "lista")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1072
  apply simp_all
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1073
 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1074
   apply (metis append_Cons append_Nil good1_flatten rflts.simps(2) rsimp.simps(2) rsimp_ALTs.elims)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1075
   by (metis (no_types, opaque_lifting) append_Cons append_Nil good1_flatten rflts.simps(2) rsimp.simps(2) rsimp_ALTs.elims)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1076
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1077
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1078
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1079
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1080
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1081
lemma simp_flatten:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1082
  shows "rsimp (RALTS ((RALTS rsa) # rsb)) = rsimp (RALTS (rsa @ rsb))"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1083
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1084
  apply(subst flatten_rsimpalts)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1085
  apply(simp add: flts_append)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1086
  by (metis Diff_empty distinct_once_enough flts_append nonalt0_fltseq nonalt_flts_rd qqq1 rdistinct_set_equality1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1087
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1088
lemma basic_rsimp_SEQ_property1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1089
  shows "rsimp_SEQ RONE r = r"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1090
  by (simp add: idiot)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1091
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1092
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1093
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1094
lemma basic_rsimp_SEQ_property3:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1095
  shows "rsimp_SEQ r RZERO = RZERO"  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1096
  using rsimp_SEQ.elims by blast
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1097
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1098
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1099
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1100
fun vsuf :: "char list \<Rightarrow> rrexp \<Rightarrow> char list list" where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1101
"vsuf [] _ = []"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1102
|"vsuf (c#cs) r1 = (if (rnullable r1) then  (vsuf cs (rder c r1)) @ [c # cs]
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1103
                                      else  (vsuf cs (rder c r1))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1104
                   ) "
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1105
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1106
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1107
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1108
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1109
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1110
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1111
fun star_update :: "char \<Rightarrow> rrexp \<Rightarrow> char list list \<Rightarrow> char list list" where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1112
"star_update c r [] = []"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1113
|"star_update c r (s # Ss) = (if (rnullable (rders r s)) 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1114
                                then (s@[c]) # [c] # (star_update c r Ss) 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1115
                               else   (s@[c]) # (star_update c r Ss) )"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1116
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1117
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1118
fun star_updates :: "char list \<Rightarrow> rrexp \<Rightarrow> char list list \<Rightarrow> char list list"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1119
  where
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1120
"star_updates [] r Ss = Ss"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1121
| "star_updates (c # cs) r Ss = star_updates cs r (star_update c r Ss)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1122
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1123
lemma stupdates_append: shows 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1124
"star_updates (s @ [c]) r Ss = star_update c r (star_updates s r Ss)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1125
  apply(induct s arbitrary: Ss)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1126
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1127
  apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1128
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1129
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1130
lemma flts_removes0:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1131
  shows "  rflts (rs @ [RZERO])  =
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1132
           rflts rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1133
  apply(induct rs)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1134
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1135
  by (metis append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1136
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1137
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1138
lemma rflts_spills_last:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1139
  shows "rflts (rs1 @ [RALTS rs]) = rflts rs1 @ rs"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1140
  apply (induct rs1 rule: rflts.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1141
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1142
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1143
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1144
lemma flts_keeps1:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1145
  shows "rflts (rs @ [RONE]) = rflts rs @ [RONE]"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1146
  apply (induct rs rule: rflts.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1147
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1148
  done
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1149
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1150
lemma flts_keeps_others:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1151
  shows "\<lbrakk>a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1\<rbrakk> \<Longrightarrow>rflts (rs @ [a]) = rflts rs @ [a]"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1152
  apply(induct rs rule: rflts.induct)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1153
  apply(auto)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1154
  by (meson k0b nonalt.elims(3))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1155
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1156
lemma spilled_alts_contained:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1157
  shows "\<lbrakk>a = RALTS rs ; a \<in> set rs1\<rbrakk> \<Longrightarrow> \<forall>r \<in> set rs. r \<in> set (rflts rs1)"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1158
  apply(induct rs1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1159
   apply simp 
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1160
  apply(case_tac "a = aa")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1161
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1162
  apply(subgoal_tac " a \<in> set rs1")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1163
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1164
   apply (meson set_ConsD)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1165
  apply(case_tac aa)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1166
  using rflts.simps(2) apply presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1167
      apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1168
  apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1169
  apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1170
  apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1171
  apply fastforce
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1172
  by simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1173
  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1174
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1175
lemma distinct_removes_duplicate_flts:
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1176
  shows " a \<in> set rsa
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1177
       \<Longrightarrow> rdistinct (rflts (map rsimp rsa @ [rsimp a])) {} =
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1178
           rdistinct (rflts (map rsimp rsa)) {}"
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1179
  apply(subgoal_tac "rsimp a \<in> set (map rsimp rsa)")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1180
  prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1181
   apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1182
  apply(induct "rsimp a")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1183
       apply simp
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1184
  using flts_removes0 apply presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1185
      apply(subgoal_tac " rdistinct (rflts (map rsimp rsa @ [rsimp a])) {} =  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1186
                          rdistinct (rflts (map rsimp rsa @ [RONE])) {}")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1187
      apply (simp only:)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1188
        apply(subst flts_keeps1)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1189
  apply (metis distinct_removes_last(1) flts_append in_set_conv_decomp rflts.simps(4))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1190
      apply presburger
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1191
        apply(subgoal_tac " rdistinct (rflts (map rsimp rsa @ [rsimp a]))    {} =  
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1192
                            rdistinct ((rflts (map rsimp rsa)) @ [RCHAR x]) {}")
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1193
      apply (simp only:)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1194
       prefer 2
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1195
       apply (metis flts_append rflts.simps(1) rflts.simps(5))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1196
  apply (metis distinct_removes_last(1) rflts_def_idiot2 rrexp.distinct(25) rrexp.distinct(3))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1197
  apply (metis distinct_removes_last(1) flts_append rflts.simps(1) rflts.simps(6) rflts_def_idiot2 rrexp.distinct(31) rrexp.distinct(5))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1198
  apply (metis distinct_removes_list rflts_spills_last spilled_alts_contained)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1199
  apply (metis distinct_removes_last(1) flts_append good.simps(1) good.simps(44) rflts.simps(1) rflts.simps(7) rflts_def_idiot2 rrexp.distinct(37))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1200
  by (metis distinct_removes_last(1) flts_append rflts.simps(1) rflts.simps(8) rflts_def_idiot2 rrexp.distinct(11) rrexp.distinct(39))
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1201
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1202
(*some basic facts about rsimp*)
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1203
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1204
unused_thms
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1205
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1206
3198605ac648 bsimp idempotency
Chengsong
parents:
diff changeset
  1207
end