thys2/SizeBound6CT.thy
author Chengsong
Tue, 01 Mar 2022 11:14:17 +0000
changeset 435 65e786a58365
parent 434 0cce1aee0fb2
child 437 43b87bab0dac
child 438 a73b2e553804
permissions -rw-r--r--
hi


theory SizeBound6CT
  imports "Lexer" "PDerivs" 
begin

section \<open>Bit-Encodings\<close>

fun orderedSufAux :: "nat \<Rightarrow> char list \<Rightarrow> char list list"
  where
 "orderedSufAux (Suc i) ss = (drop i ss) # (orderedSufAux i ss)"
|"orderedSufAux 0 ss = Nil"

fun 
orderedSuf :: "char list \<Rightarrow> char list list"
where
"orderedSuf s = orderedSufAux (length s) s"

fun orderedPrefAux :: "nat \<Rightarrow> char list \<Rightarrow> char list list"
  where
"orderedPrefAux (Suc i) ss = (take i ss) # (orderedPrefAux i ss)"
|"orderedPrefAux 0 ss = Nil"


fun orderedPref :: "char list \<Rightarrow> char list list"
  where
"orderedPref s = orderedPrefAux (length s) s"

lemma shape_of_pref_1list:
  shows "orderedPref [c] = [[]]"
  apply auto
  done

lemma shape_of_suf_1list:
  shows "orderedSuf [c] = [[c]]"
  by auto

lemma shape_of_suf_2list:
  shows "orderedSuf [c2, c3] = [[c3], [c2,c3]]"
  by auto

lemma shape_of_prf_2list:
  shows "orderedPref [c1, c2] = [[c1], []]"
  by auto


lemma shape_of_suf_3list:
  shows "orderedSuf [c1, c2, c3] = [[c3], [c2, c3], [c1, c2, c3]]"
  by auto
(*
lemma throwing_elem_around:
  shows "orderedSuf (s1 @ [a] @ s) = (orderedSuf s) @ (map (\<lambda>s11. s11 @ s) (orderedSuf ( s1 @ [a]) ))"
and "orderedSuf (s1 @ [a] @ s) = (orderedSuf ([a] @ s) @ (map (\<lambda>s11. s11 @ ([a] @ s))) (orderedSuf s1) )"
  sorry


lemma suf_cons:
  shows "orderedSuf (s1 @ s) = (orderedSuf s) @ (map (\<lambda>s11. s11 @ s) (orderedSuf s1))"
  apply(induct s arbitrary: s1)
   apply simp
  apply(subgoal_tac "s1 @ a # s = (s1 @ [a]) @ s")
  prefer 2
   apply simp
  apply(subgoal_tac "orderedSuf (s1 @ a # s) = orderedSuf ((s1 @ [a]) @ s)")
  prefer 2
   apply presburger
  apply(drule_tac x="s1 @ [a]" in meta_spec)
  sorry
*)


lemma shape_of_prf_3list:
  shows "orderedPref [c1, c2, c3] = [[c1, c2], [c1], []]"
  by auto

fun zip_concat :: "char list list \<Rightarrow> char list list \<Rightarrow> char list list"
  where 
    "zip_concat (s1#ss1) (s2#ss2) = (s1@s2) # (zip_concat ss1 ss2)"
  |   "zip_concat [] [] = []"
  | "zip_concat [] (s2#ss2) = s2 # (zip_concat [] ss2)"
  | "zip_concat (s1#ss1) [] = s1 # (zip_concat ss1 [])"


(*
lemma compliment_pref_suf:
  shows "zip_concat (orderedPref s) (orderedSuf s) = replicate (length s) s "
  apply(induct s)
   apply auto[1]
  sorry


*)

datatype rrexp = 
  RZERO
| RONE 
| RCHAR char
| RSEQ rrexp rrexp
| RALTS "rrexp list"
| RSTAR rrexp

abbreviation
  "RALT r1 r2 \<equiv> RALTS [r1, r2]"



fun
 rnullable :: "rrexp \<Rightarrow> bool"
where
  "rnullable (RZERO) = False"
| "rnullable (RONE  ) = True"
| "rnullable (RCHAR   c) = False"
| "rnullable (RALTS   rs) = (\<exists>r \<in> set rs. rnullable r)"
| "rnullable (RSEQ  r1 r2) = (rnullable r1 \<and> rnullable r2)"
| "rnullable (RSTAR   r) = True"


fun
 rder :: "char \<Rightarrow> rrexp \<Rightarrow> rrexp"
where
  "rder c (RZERO) = RZERO"
| "rder c (RONE) = RZERO"
| "rder c (RCHAR d) = (if c = d then RONE else RZERO)"
| "rder c (RALTS rs) = RALTS (map (rder c) rs)"
| "rder c (RSEQ r1 r2) = 
     (if rnullable r1
      then RALT   (RSEQ (rder c r1) r2) (rder c r2)
      else RSEQ   (rder c r1) r2)"
| "rder c (RSTAR r) = RSEQ  (rder c r) (RSTAR r)"


fun 
  rders :: "rrexp \<Rightarrow> string \<Rightarrow> rrexp"
where
  "rders r [] = r"
| "rders r (c#s) = rders (rder c r) s"

fun rdistinct :: "'a list \<Rightarrow>'a set \<Rightarrow> 'a list"
  where
  "rdistinct [] acc = []"
| "rdistinct (x#xs)  acc = 
     (if x \<in> acc then rdistinct xs  acc 
      else x # (rdistinct xs  ({x} \<union> acc)))"





fun rflts :: "rrexp list \<Rightarrow> rrexp list"
  where 
  "rflts [] = []"
| "rflts (RZERO # rs) = rflts rs"
| "rflts ((RALTS rs1) # rs) = rs1 @ rflts rs"
| "rflts (r1 # rs) = r1 # rflts rs"


fun rsimp_ALTs :: " rrexp list \<Rightarrow> rrexp"
  where
  "rsimp_ALTs  [] = RZERO"
| "rsimp_ALTs [r] =  r"
| "rsimp_ALTs rs = RALTS rs"

fun rsimp_SEQ :: " rrexp \<Rightarrow> rrexp \<Rightarrow> rrexp"
  where
  "rsimp_SEQ  RZERO _ = RZERO"
| "rsimp_SEQ  _ RZERO = RZERO"
| "rsimp_SEQ RONE r2 = r2"
| "rsimp_SEQ r1 r2 = RSEQ r1 r2"


fun rsimp :: "rrexp \<Rightarrow> rrexp" 
  where
  "rsimp (RSEQ r1 r2) = rsimp_SEQ  (rsimp r1) (rsimp r2)"
| "rsimp (RALTS rs) = rsimp_ALTs  (rdistinct (rflts (map rsimp rs))  {}) "
| "rsimp r = r"


fun 
  rders_simp :: "rrexp \<Rightarrow> string \<Rightarrow> rrexp"
where
  "rders_simp r [] = r"
| "rders_simp r (c#s) = rders_simp (rsimp (rder c r)) s"

fun rsize :: "rrexp \<Rightarrow> nat" where
  "rsize RZERO = 1"
| "rsize (RONE) = 1" 
| "rsize (RCHAR  c) = 1"
| "rsize (RALTS  rs) = Suc (sum_list (map rsize rs))"
| "rsize (RSEQ  r1 r2) = Suc (rsize r1 + rsize r2)"
| "rsize (RSTAR  r) = Suc (rsize r)"


fun rlist_size :: "rrexp list \<Rightarrow> nat" where
"rlist_size (r # rs) = rsize r + rlist_size rs" 
| "rlist_size [] = 0"

thm neq_Nil_conv
lemma hand_made_def_rlist_size:
  shows "rlist_size rs = sum_list (map rsize rs)"
proof (induct rs)
  case Nil show ?case by simp
next
  case (Cons a rs) thus ?case
    by simp
qed

lemma rsimp_aalts_smaller:
  shows "rsize (rsimp_ALTs  rs) \<le> rsize (RALTS rs)"
  apply(induct rs)
   apply simp
  apply simp
  apply(case_tac "rs = []")
   apply simp
  apply(subgoal_tac "\<exists>rsp ap. rs = ap # rsp")
   apply(erule exE)+
   apply simp
  apply simp
  by(meson neq_Nil_conv)
  





(*
lemma rders_shape:
  shows "s \<noteq> [] \<Longrightarrow> rders_simp (RSEQ r1 r2) s = 
         rsimp (RALTS  ((RSEQ (rders r1 s) r2) #
           (map (rders r2) (orderedSuf s))) )"
  apply(induct s arbitrary: r1 r2 rule: rev_induct)
   apply simp
  apply simp

  sorry
*)



fun rders_cond_list :: "rrexp \<Rightarrow> bool list \<Rightarrow> char list list \<Rightarrow> rrexp list"
  where
"rders_cond_list r2 (True # bs) (s # strs) = (rders_simp r2 s) # (rders_cond_list r2 bs strs)"
| "rders_cond_list r2 (False # bs) (s # strs) = rders_cond_list r2 bs strs"
| "rders_cond_list r2 [] s = []"
| "rders_cond_list r2 bs [] = []"

fun nullable_bools :: "rrexp \<Rightarrow> char list list \<Rightarrow> bool list"
  where
"nullable_bools r (s#strs) = (rnullable (rders r s)) # (nullable_bools r strs) "
|"nullable_bools r [] = []"

fun cond_list :: "rrexp \<Rightarrow> rrexp \<Rightarrow> char list \<Rightarrow> rrexp list"
  where
"cond_list r1 r2 s = rders_cond_list r2 (nullable_bools r1 (orderedPref s) ) (orderedSuf s)"


lemma rSEQ_mono:
  shows "rsize (rsimp_SEQ r1 r2) \<le>rsize ( RSEQ r1 r2)"
  apply auto
  apply(induct r1)
       apply auto
      apply(case_tac "r2")
       apply simp_all
     apply(case_tac r2)
          apply simp_all
     apply(case_tac r2)
         apply simp_all
     apply(case_tac r2)
        apply simp_all
     apply(case_tac r2)
  apply simp_all
  done

lemma ralts_cap_mono:
  shows "rsize (RALTS rs) \<le> Suc ( sum_list (map rsize rs)) "
  by simp

lemma rflts_def_idiot:
  shows "\<lbrakk> a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1\<rbrakk>
       \<Longrightarrow> rflts (a # rs) = a # rflts rs"
  apply(case_tac a)
       apply simp_all
  done


lemma rflts_mono:
  shows "sum_list (map rsize (rflts rs))\<le> sum_list (map rsize rs)"
  apply(induct rs)
  apply simp
  apply(case_tac "a = RZERO")
   apply simp
  apply(case_tac "\<exists>rs1. a = RALTS rs1")
  apply(erule exE)
   apply simp
  apply(subgoal_tac "rflts (a # rs) = a # (rflts rs)")
  prefer 2
  using rflts_def_idiot apply blast
  apply simp
  done

lemma rdistinct_smaller: shows "sum_list (map rsize (rdistinct rs ss)) \<le>
sum_list (map rsize rs )"
  apply (induct rs arbitrary: ss)
   apply simp
  by (simp add: trans_le_add2)

lemma rdistinct_phi_smaller: "sum_list (map rsize (rdistinct rs {})) \<le> sum_list (map rsize rs)"
  by (simp add: rdistinct_smaller)


lemma rsimp_alts_mono :
  shows "\<And>x. (\<And>xa. xa \<in> set x \<Longrightarrow> rsize (rsimp xa) \<le> rsize xa)  \<Longrightarrow>
rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {})) \<le> Suc (sum_list (map rsize x))"
  apply(subgoal_tac "rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {} )) 
                    \<le> rsize (RALTS (rdistinct (rflts (map rsimp x)) {} ))")
  prefer 2
  using rsimp_aalts_smaller apply auto[1]
  apply(subgoal_tac "rsize (RALTS (rdistinct (rflts (map rsimp x)) {})) \<le>Suc( sum_list (map rsize (rdistinct (rflts (map rsimp x)) {})))")
  prefer 2
  using ralts_cap_mono apply blast
  apply(subgoal_tac "sum_list (map rsize (rdistinct (rflts (map rsimp x)) {})) \<le>
                     sum_list (map rsize ( (rflts (map rsimp x))))")
  prefer 2
  using rdistinct_smaller apply presburger
  apply(subgoal_tac "sum_list (map rsize (rflts (map rsimp x))) \<le> 
                     sum_list (map rsize (map rsimp x))")
  prefer 2
  using rflts_mono apply blast
  apply(subgoal_tac "sum_list (map rsize (map rsimp x)) \<le> sum_list (map rsize x)")
  prefer 2
  
  apply (simp add: sum_list_mono)
  by linarith





lemma rsimp_mono:
  shows "rsize (rsimp r) \<le> rsize r"

  apply(induct r)
             apply simp_all

  apply(subgoal_tac "rsize (rsimp_SEQ (rsimp r1) (rsimp r2)) \<le> rsize (RSEQ (rsimp r1) (rsimp r2))")
  
    apply force
  using rSEQ_mono
   apply presburger
  using rsimp_alts_mono by auto

lemma idiot:
  shows "rsimp_SEQ RONE r = r"
  apply(case_tac r)
       apply simp_all
  done

lemma no_alt_short_list_after_simp:
  shows "RALTS rs = rsimp r \<Longrightarrow> rsimp_ALTs rs = RALTS rs"
  sorry

lemma no_further_dB_after_simp:
  shows "RALTS rs = rsimp r \<Longrightarrow> rdistinct rs {} = rs"

  sorry


lemma idiot2:
  shows " \<lbrakk>r1 \<noteq> RZERO; r1 \<noteq> RONE;r2 \<noteq> RZERO\<rbrakk>
    \<Longrightarrow> rsimp_SEQ r1 r2 = RSEQ r1 r2"
  apply(case_tac r1)
       apply(case_tac r2)
  apply simp_all
     apply(case_tac r2)
  apply simp_all
     apply(case_tac r2)
  apply simp_all
   apply(case_tac r2)
  apply simp_all
  apply(case_tac r2)
       apply simp_all
  done

lemma rsimp_aalts_another:
  shows "\<forall>r \<in> (set  (map rsimp  ((RSEQ (rders r1 s) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))  )) ). (rsize r) < N "
  sorry



lemma shape_derssimpseq_onechar:
  shows "   (rders_simp r [c]) =  (rsimp (rders r [c]))"
and "rders_simp (RSEQ r1 r2) [c] = 
         rsimp (RALTS  ((RSEQ (rders r1 [c]) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref [c]))  (orderedSuf [c]))) )"
   apply simp
  apply(simp add: rders.simps)
  apply(case_tac "rsimp (rder c r1) = RZERO")
   apply auto
  apply(case_tac "rsimp (rder c r1) = RONE")
   apply auto
   apply(subgoal_tac "rsimp_SEQ RONE (rsimp r2) = rsimp r2")
  prefer 2
  using idiot
    apply simp
   apply(subgoal_tac "rsimp_SEQ RONE (rsimp r2) = rsimp_ALTs (rdistinct (rflts [rsimp r2]) {})")
    prefer 2
    apply auto
   apply(case_tac "rsimp r2")
        apply auto
   apply(subgoal_tac "rdistinct x5 {} = x5")
  prefer 2
  using no_further_dB_after_simp
    apply metis
   apply(subgoal_tac "rsimp_ALTs (rdistinct x5 {}) = rsimp_ALTs x5")
    prefer 2
    apply fastforce  
   apply auto
   apply (metis no_alt_short_list_after_simp)
  apply (case_tac "rsimp r2 = RZERO")
   apply(subgoal_tac "rsimp_SEQ (rsimp (rder c r1)) (rsimp r2) = RZERO")
    prefer 2
    apply(case_tac "rsimp ( rder c r1)")
         apply auto
  apply(subgoal_tac "rsimp_SEQ (rsimp (rder c r1)) (rsimp r2) = RSEQ (rsimp (rder c r1)) (rsimp r2)")
   prefer 2
   apply auto
  sorry



lemma shape_derssimpseq_onechar2:
  shows "rders_simp (RSEQ r1 r2) [c] = 
         rsimp (RALTS  ((RSEQ (rders_simp r1 [c]) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref [c]))  (orderedSuf [c]))) )"
  sorry


lemma rders__onechar:
  shows " (rders_simp r [c]) =  (rsimp (rders r [c]))"
  by simp

lemma rders_append:
  "rders c (s1 @ s2) = rders (rders c s1) s2"
  apply(induct s1 arbitrary: c s2)
  apply(simp_all)
  done

lemma rders_simp_append:
  "rders_simp c (s1 @ s2) = rders_simp (rders_simp c s1) s2"
  apply(induct s1 arbitrary: c s2)
  apply(simp_all)
  done

lemma inside_simp_removal:
  shows " rsimp (rder x (rsimp r)) = rsimp (rder x r)"
  
  sorry

lemma set_related_list:
  shows "distinct rs  \<Longrightarrow> length rs = card (set rs)"
  by (simp add: distinct_card)
(*this section deals with the property of distinctBy: creates a list without duplicates*)
lemma rdistinct_never_added_twice:
  shows "rdistinct (a # rs) {a} = rdistinct rs {a}"
  by force


lemma rdistinct_does_the_job:
  shows "distinct (rdistinct rs s)"
  apply(induct rs arbitrary: s)
   apply simp
  apply simp
  sorry


lemma simp_helps_der_pierce:
  shows " rsimp
            (rder x
              (rsimp_ALTs rs)) = 
          rsimp 
            (rsimp_ALTs 
              (map (rder x )
                rs
              )
            )"
  sorry

lemma simp_helps_der_pierce_dB:
  shows " rsimp
            (rsimp_ALTs
              (map (rder x)
                (rdistinct rs {}))) = 
          rsimp (rsimp_ALTs (rdistinct (map (rder x) rs) {}))"

  sorry

lemma simp_helps_der_pierce_flts:
  shows " rsimp
            (rsimp_ALTs
             (rdistinct 
                (map (rder x)
                  (rflts rs )
                ) {}
             )
            ) = 
          rsimp (rsimp_ALTs (rdistinct (rflts (map (rder x) rs)) {})  )"

  sorry


lemma unfold_ders_simp_inside_only: 
  shows "    (rders_simp (RSEQ r1 r2) xs =
           rsimp (RALTS (RSEQ (rders_simp r1 xs) r2 # rders_cond_list r2 (nullable_bools r1 (orderedPref xs)) (orderedSuf xs))))
       \<Longrightarrow> rsimp (rder x (rders_simp (RSEQ r1 r2) xs)) = rsimp (rder x (rsimp (RALTS(RSEQ (rders_simp r1 xs) r2 # rders_cond_list r2 (nullable_bools r1 (orderedPref xs)) (orderedSuf xs)))))"
 by presburger



lemma unfold_ders_simp_inside_only_nosimp: 
  shows "    (rders_simp (RSEQ r1 r2) xs =
           rsimp (RALTS (RSEQ (rders_simp r1 xs) r2 # rders_cond_list r2 (nullable_bools r1 (orderedPref xs)) (orderedSuf xs))))
       \<Longrightarrow> rsimp (rder x (rders_simp (RSEQ r1 r2) xs)) = rsimp (rder x (RALTS(RSEQ (rders_simp r1 xs) r2 # rders_cond_list r2 (nullable_bools r1 (orderedPref xs)) (orderedSuf xs))))"
  using inside_simp_removal by presburger




lemma rders_simp_one_char:
  shows "rders_simp r [c] = rsimp (rder c r)"
  apply auto
  done

lemma rsimp_idem:
  shows "rsimp (rsimp r) = rsimp r"
  sorry

corollary rsimp_inner_idem1:
  shows "rsimp r = RSEQ r1 r2 \<Longrightarrow> rsimp r1 = r1 \<and> rsimp r2 = r2"
  
  sorry

corollary rsimp_inner_idem2:
  shows "rsimp r = RALTS rs \<Longrightarrow> \<forall>r' \<in> (set rs). rsimp r' = r'"
  sorry

corollary rsimp_inner_idem3:
  shows "rsimp r = RALTS rs \<Longrightarrow> map rsimp rs = rs"
  by (meson map_idI rsimp_inner_idem2)

corollary rsimp_inner_idem4:
  shows "rsimp r = RALTS rs \<Longrightarrow> flts rs = rs"
  sorry


lemma head_one_more_simp:
  shows "map rsimp (r # rs) = map rsimp (( rsimp r) # rs)"
  by (simp add: rsimp_idem)

lemma head_one_more_dersimp:
  shows "map rsimp ((rder x (rders_simp r s) # rs)) = map rsimp ((rders_simp r (s@[x]) ) # rs)"
  using head_one_more_simp rders_simp_append rders_simp_one_char by presburger

thm cond_list.simps

lemma suffix_plus1char:
  shows "\<not> (rnullable (rders r1 s)) \<Longrightarrow> cond_list r1 r2 (s@[c]) = map (rder c) (cond_list r1 r2 s)"
   apply simp
  sorry

lemma suffix_plus1charn:
shows "rnullable (rders r1 s) \<Longrightarrow> cond_list r1 r2 (s@[c]) = (rder c r2) # (map (rder c) (cond_list r1 r2 s))"
  sorry

lemma ders_simp_nullability:
  shows "rnullable (rders r s) = rnullable (rders_simp r s)"
  sorry

lemma  first_elem_seqder:
  shows "\<not>rnullable r1p \<Longrightarrow> map rsimp (rder x (RSEQ r1p r2)
                   # rs) = map rsimp ((RSEQ (rder x r1p) r2) # rs) "
  by auto

lemma first_elem_seqder1:
  shows  "\<not>rnullable (rders_simp r xs) \<Longrightarrow> map rsimp ( (rder x (RSEQ (rders_simp r xs) r2)) # rs) = 
                                          map rsimp ( (RSEQ (rsimp (rder x (rders_simp r xs))) r2) # rs)"
  by (simp add: rsimp_idem)

lemma first_elem_seqdersimps:
  shows "\<not>rnullable (rders_simp r xs) \<Longrightarrow> map rsimp ( (rder x (RSEQ (rders_simp r xs) r2)) # rs) = 
                                          map rsimp ( (RSEQ (rders_simp r (xs @ [x])) r2) # rs)"
  using first_elem_seqder1 rders_simp_append by auto

lemma first_elem_seqder_nullable:
  shows "rnullable (rders_simp r1 xs) \<Longrightarrow> cond_list r1 r2 (xs @ [x]) = rder x r2 # map (rder x) (cond_list r1 r2 xs)"
  sorry


(*nullable_seq_with_list1 related *)
lemma LHS0_def_der_alt:
  shows "rsimp (rder x (RALTS ( (RSEQ (rders_simp r1 xs) r2) # (cond_list r1 r2 xs)))) = 
         rsimp (RALTS ((rder x (RSEQ (rders_simp r1 xs) r2)) # (map (rder x) (cond_list r1 r2 xs))))"
  by fastforce

lemma LHS1_def_der_seq:
  shows "rnullable (rders_simp r1 xs) \<Longrightarrow> 
rsimp (rder x (RALTS ( (RSEQ (rders_simp r1 xs) r2) # (cond_list r1 r2 xs)))) = 
rsimp(RALTS ((RALTS ((RSEQ (rders_simp r1 (xs @ [x])) r2) # [rder x r2]) ) # (map (rder x ) (cond_list r1 r2 xs))))"
  by (simp add: rders_simp_append rsimp_idem)





lemma cond_list_head_last:
  shows "rnullable (rders r1 s) \<Longrightarrow> 
        RALTS (r # (cond_list r1 r2 (s @ [c]))) = RALTS (r # ((rder c r2) # (map (rder c) (cond_list r1 r2 s))))"
  using suffix_plus1charn by blast

lemma simp_flatten2:
  shows "rsimp (RALTS (r # [RALTS rs])) = rsimp (RALTS (r # rs))"
  sorry


lemma simp_flatten:
  shows "rsimp (RALTS ((RALTS rsa) # rsb)) = rsimp (RALTS (rsa @ rsb))"

  sorry

lemma RHS1:
  shows "rnullable (rders_simp r1 xs) \<Longrightarrow>
                                    rsimp (RALTS ((RSEQ (rders_simp r1 (xs @ [x])) r2) # 
(cond_list r1 r2 (xs @[x])))) = 
                                    rsimp (RALTS ((RSEQ (rders_simp r1 (xs @ [x])) r2) # 
( ((rder x r2) # (map (rder x) (cond_list r1 r2 xs)))))) "
  using first_elem_seqder_nullable by presburger


lemma nullable_seq_with_list1:
  shows " rnullable (rders_simp r1 s) \<Longrightarrow>
    rsimp (rder c (RALTS ( (RSEQ (rders_simp r1 s) r2) #  (cond_list r1 r2 s)) )) =
    rsimp (RALTS ( (RSEQ (rders_simp r1 (s @ [c])) r2) # (cond_list r1 r2 (s @ [c])) ) )"
  using RHS1 LHS1_def_der_seq cond_list_head_last simp_flatten
  by (metis append.left_neutral append_Cons)


(*^^^^^^^^^nullable_seq_with_list1 related ^^^^^^^^^^^^^^^^*)










lemma nullable_seq_with_list:
  shows "rnullable (rders_simp r1 xs) \<Longrightarrow> rsimp (rder x (RALTS ((RSEQ (rders_simp r1 xs) r2) # 
      (rders_cond_list r2 (nullable_bools r1 (orderedPref xs)) (orderedSuf xs)) ))) = 
                                          rsimp (RALTS ((RSEQ (rders_simp r1 (xs@[x])) r2) #
      (rders_cond_list r2 (nullable_bools r1 (orderedPref (xs@[x]))) (orderedSuf (xs@[x]))) ) )
    "
  apply(subgoal_tac "rsimp (rder x (RALTS ((RSEQ (rders_simp r1 xs) r2) # (cond_list r1 r2 xs)))) = 
                     rsimp (RALTS ((RSEQ (rders_simp r1 (xs@[x])) r2) # (cond_list r1 r2 (xs@[x]))))")
   apply auto[1]
  using nullable_seq_with_list1 by auto




lemma r1r2ders_whole:
"rsimp
         (RALTS
           (rder x (RSEQ (rders_simp r1 xs) r2) #
            map (rder x) (rders_cond_list r2 (nullable_bools r1 (orderedPref xs)) (orderedSuf xs)))) = 
      rsimp(    RALTS(   ( (RSEQ (rders_simp r1 (xs@[x])) r2)
                   #  (rders_cond_list r2 (nullable_bools r1 (orderedPref (xs @ [x]))) (orderedSuf (xs @ [x])))))) "
  using ders_simp_nullability first_elem_seqdersimps nullable_seq_with_list1 suffix_plus1char by auto

lemma rders_simp_same_simpders:
  shows "s \<noteq> [] \<Longrightarrow> rders_simp r s = rsimp (rders r s)"
  apply(induct s rule: rev_induct)
   apply simp
  apply(case_tac "xs = []")
   apply simp
  apply(simp add: rders_append rders_simp_append)
  using inside_simp_removal by blast

lemma shape_derssimp_seq:
  shows "\<lbrakk>s \<noteq> []\<rbrakk> \<Longrightarrow> rders_simp (RSEQ r1 r2) s = 
         rsimp (RALTS  ((RSEQ (rders_simp r1 s) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))) )"

  apply(induct s arbitrary: r1 r2 rule: rev_induct)
   apply simp
  apply(case_tac "xs = []")
  using shape_derssimpseq_onechar2 apply force
  apply(simp only: rders_simp_append)
  apply(simp only: rders_simp_one_char)

  apply(subgoal_tac "rsimp (rder x (rders_simp (RSEQ r1 r2) xs))
                   = rsimp (rder x (RALTS(RSEQ (rders_simp r1 xs) r2 # rders_cond_list r2 (nullable_bools r1 (orderedPref xs)) (orderedSuf xs))))")
   prefer 2
  using unfold_ders_simp_inside_only_nosimp apply presburger
  apply(subgoal_tac "rsimp (rder x (RALTS (RSEQ (rders_simp r1 xs) r2
                   # rders_cond_list r2 (nullable_bools r1 (orderedPref xs)) (orderedSuf xs)))) = 
                     rsimp ( (RALTS (rder x (RSEQ (rders_simp r1 xs) r2)
                   # (map (rder x) (rders_cond_list r2 (nullable_bools r1 (orderedPref xs)) (orderedSuf xs))))))
                      ")
   prefer 2
   apply simp
  using r1r2ders_whole rders_simp_append rders_simp_one_char by presburger

(*

  apply(subgoal_tac " rsimp
            (rder x
              (rsimp_ALTs
                (rdistinct
                  (rflts
                    (rsimp_SEQ (rsimp (rders_simp r1 xs)) (rsimp r2) #
                     map rsimp
                      (rders_cond_list r2 (nullable_bools r1 (orderedPrefAux (length xs) xs)) (orderedSufAux (length xs) xs))))
                  {}))) =  
                      rsimp
            (
              (rsimp_ALTs
               (map (rder x)
                (rdistinct
                  (rflts
                    (rsimp_SEQ (rsimp (rders_simp r1 xs)) (rsimp r2) #
                     map rsimp
                      (rders_cond_list r2 (nullable_bools r1 (orderedPrefAux (length xs) xs)) (orderedSufAux (length xs) xs))))
                  {})
               )
              )
            ) ")
   prefer 2
*)

lemma shape_derssimp_alts:
  shows "s \<noteq> [] \<Longrightarrow> rders_simp (RALTS rs) s = rsimp (RALTS (map (\<lambda>r. rders r s) rs))"
  apply(case_tac "s")
   apply simp
  apply simp
  sorry
(*
fun rexp_encode :: "rrexp \<Rightarrow> nat"
  where
"rexp_encode RZERO = 0"
|"rexp_encode RONE = 1"
|"rexp_encode (RCHAR c) = 2"
|"rexp_encode (RSEQ r1 r2) = ( 2 ^ (rexp_encode r1)) "
*)
lemma finite_chars:
  shows " \<exists>N. ( (\<forall>r \<in> (set rs). \<exists>c. r = RCHAR c) \<and> (distinct rs) \<longrightarrow> length rs < N)"
  apply(rule_tac x = "Suc 256" in exI)
  sorry

definition all_chars :: "int \<Rightarrow> char list"
  where "all_chars n = map char_of [0..n]"
(*
fun rexp_enum :: "nat \<Rightarrow> rrexp list"
  where 
"rexp_enum 0 = []"
|"rexp_enum (Suc 0) =  RALTS [] # RZERO # (RONE # (map RCHAR (all_chars 255)))"
|"rexp_enum (Suc n) = [(RSEQ r1 r2). r1 \<in> set (rexp_enum i) \<and> r2 \<in> set (rexp_enum j) \<and> i + j = n]"

*)

fun rexp_enum :: "nat \<Rightarrow> rrexp set"
  where 
"rexp_enum 0 = {}"
|"rexp_enum (Suc 0) =  {RALTS []} \<union> {RZERO} \<union> {RONE} \<union> { (RCHAR c) |c. c \<in> set (all_chars 255)}"
|"rexp_enum (Suc n) = {(RSEQ r1 r2)|r1 r2 i j. r1 \<in>  (rexp_enum i) \<and> r2 \<in>  (rexp_enum j) \<and> i + j = n}"

 
lemma finite_sized_rexp_forms_finite_set:
  shows " \<exists>SN. ( \<forall>r.( rsize r < N \<longrightarrow> r \<in> SN)) \<and> (finite SN)"
  apply(induct N)
   apply simp
   apply auto
 (*\<lbrakk>\<forall>r. rsize r < N \<longrightarrow> r \<in> SN; finite SN\<rbrakk> \<Longrightarrow> \<exists>SN. (\<forall>r. rsize r < Suc N \<longrightarrow> r \<in> SN) \<and> finite SN*)
 (* \<And>N. \<exists>SN. (\<forall>r. rsize r < N \<longrightarrow> r \<in> SN) \<and> finite SN \<Longrightarrow> \<exists>SN. (\<forall>r. rsize r < Suc N \<longrightarrow> r \<in> SN) \<and> finite SN*)
  sorry


lemma finite_size_finite_regx:
  shows " \<exists>l. \<forall>rs. ((\<forall>r \<in> (set rs). rsize r < N) \<and> (distinct rs) \<longrightarrow> (length rs) < l) "
  sorry

(*below  probably needs proved concurrently*)

lemma finite_r1r2_ders_list:
  shows "(\<forall>s. rsize (rders_simp r1 s) < N1 \<and> rsize (rders_simp r2 s) < N2)
           \<Longrightarrow>  \<exists>l. \<forall>s.
(length (rdistinct  (map rsimp (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))) {}) )  < l "
  sorry

(*
\<lbrakk>s \<noteq> []\<rbrakk> \<Longrightarrow> rders_simp (RSEQ r1 r2) s = 
         rsimp (RALTS  ((RSEQ (rders r1 s) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))) )
*)
 

lemma sum_list_size2:
  shows "\<forall>z \<in>set rs. (rsize z ) \<le> Nr \<Longrightarrow> rlist_size rs \<le> (length rs) * Nr"
  apply(induct rs)
   apply simp
  by simp

lemma sum_list_size:
  fixes rs
  shows " \<forall>r \<in> (set rs). (rsize r) \<le> Nr \<and> (length rs) \<le> l \<Longrightarrow> rlist_size rs \<le> l * Nr"
  by (metis dual_order.trans mult.commute mult_le_mono2 mult_zero_right sum_list_size2 zero_le)

lemma seq_second_term_chain1:
  shows " \<forall>s. rlist_size (rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s)) \<le> 
 rlist_size (rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s))"

  sorry


lemma seq_second_term_chain2:
  shows "\<forall>s.  rlist_size (rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s)) = 
 rlist_size (map rsimp (rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s)))"

  oops

lemma seq_second_term_bounded:
  fixes r2 r1
  assumes "\<forall>s. rsize (rders_simp r2 s) < N2"
  shows "\<exists>N3. \<forall>s. rlist_size (rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s)) < N3"

  sorry


lemma seq_first_term_bounded:
  fixes r1 r2
  shows "\<exists>Nr. \<forall>s. rsize (rders_simp r1 s) \<le> Nr \<Longrightarrow> \<exists>Nr'. \<forall>s. rsize (RSEQ  (rders_simp r1 s) r2) \<le> Nr'"
  apply(erule exE)
  apply(rule_tac x = "Nr + (rsize r2) + 1" in exI)
  by simp


lemma alts_triangle_inequality:
  shows "rsize (RALTS (r # rs)) \<le> rsize r + rlist_size rs + 1 "
  apply(subgoal_tac "rsize (RALTS (r # rs) ) =  rsize r + rlist_size rs + 1")
   apply auto[1]
  apply(induct rs)
   apply simp
  apply auto
  done

lemma seq_equal_term_nosimp_entire_bounded:
  shows "(\<forall>s. rsize (rders_simp r1 s) < N1 \<and> rsize (rders_simp r2 s) < N2) 
      \<Longrightarrow> \<exists>N3. \<forall>s.  rsize  (RALTS  ((RSEQ (rders_simp r1 s) r2) #
          (rdistinct ((rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s)) ) {}) ) ) \<le> N3"
  apply(subgoal_tac "\<exists>N3. \<forall>s. rsize  (RALTS  ((RSEQ (rders_simp r1 s) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))) ) \<le>
                                        rsize (RSEQ (rders_simp r1 s) r2) + 
            rlist_size (map rsimp (rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s))) + 1")
  prefer 2
  using alts_triangle_inequality apply presburger
  using seq_first_term_bounded
  using seq_second_term_bounded
  apply(subgoal_tac "\<exists>N3. \<forall>s. rlist_size (rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s)) < N3")
  prefer 2
   apply meson
  apply(subgoal_tac "\<exists>Nr'. \<forall>s. rsize (RSEQ  (rders_simp r1 s) r2) \<le> Nr'")
   prefer 2
  apply (meson order_le_less)
  apply(erule exE)
  apply(erule exE)
  sorry

lemma alts_simp_bounded_by_sloppy1_version:
  shows "\<forall>s. rsize (rsimp (RALTS  ((RSEQ (rders_simp r1 s) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))) )) \<le> 
rsize  (RALTS (rdistinct ((RSEQ (rders_simp r1 s) r2) #
                          (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))
                         ) 
                         {}
              )
        ) "
  sorry

lemma alts_simp_bounded_by_sloppy1:
  shows "rsize  (rsimp (RALTS (rdistinct ((RSEQ (rders_simp r1 s) r2) #
                          (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))
                         ) 
                         {}
              )
        )) \<le>
rsize  (RALTS  ((RSEQ (rders_simp r1 s) r2) #
                          (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))
                         ) 
        )"
  sorry



(*this section deals with the property of distinctBy: creates a list without duplicates*)
lemma distinct_mono:
  shows "rlist_size (rdistinct (a # s) {}) \<le> rlist_size (a # (rdistinct s {}) )"
  sorry

lemma distinct_acc_mono:
  shows "A \<subseteq> B \<Longrightarrow> rlist_size (rdistinct s A) \<ge> rlist_size (rdistinct s B)"
  apply(induct s arbitrary: A B)
   apply simp
  apply(case_tac "a \<in> A")
  apply(subgoal_tac "a \<in> B")
  
  apply simp
  
   apply blast
  apply(subgoal_tac "rlist_size (rdistinct (a # s) A) = rlist_size (a # (rdistinct s (A \<union> {a})))")
   apply(case_tac "a \<in> B")
  apply(subgoal_tac "rlist_size (rdistinct (a # s) B) = rlist_size ( (rdistinct s B))") 
  apply (metis Un_insert_right dual_order.trans insert_subset le_add_same_cancel2 rlist_size.simps(1) sup_bot_right zero_order(1))
  apply simp
  apply auto
  by (meson insert_mono)


lemma distinct_mono2:
  shows " rlist_size (rdistinct s {a}) \<le> rlist_size (rdistinct s {})"
  apply(induct s)
   apply simp
  apply simp
  using distinct_acc_mono by auto



lemma distinct_mono_spares_first_elem:
  shows "rsize (RALTS (rdistinct (a # s) {})) \<le> rsize (RALTS (a # (rdistinct s {})))"
  apply simp
  apply (subgoal_tac "rlist_size ( (rdistinct s {a})) \<le> rlist_size ( (rdistinct s {})) ")
  using hand_made_def_rlist_size apply auto[1]
  using distinct_mono2 by auto

lemma sloppy1_bounded_by_sloppiest:
  shows "rsize  (RALTS (rdistinct ((RSEQ (rders_simp r1 s) r2) #
                          (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))
                         ) 
                         {}
              )
        ) \<le> rsize  (RALTS ((RSEQ (rders_simp r1 s) r2) #
                          (rdistinct (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s)) {})
                         
                         
              )
        )"
  
  sorry


lemma alts_simp_bounded_by_sloppiest_version:
  shows "\<forall>s. rsize (rsimp (RALTS  ((RSEQ (rders_simp r1 s) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))) )) \<le> 
rsize  (RALTS  ((RSEQ (rders_simp r1 s) r2) #
           (rdistinct (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s)) {})  ) ) "
  by (meson alts_simp_bounded_by_sloppy1_version order_trans sloppy1_bounded_by_sloppiest)


lemma seq_equal_term_entire_bounded:
  shows "(\<forall>s. rsize (rders_simp r1 s) < N1 \<and> rsize (rders_simp r2 s) < N2) 
      \<Longrightarrow> \<exists>N3. \<forall>s.  rsize (rsimp (RALTS  ((RSEQ (rders_simp r1 s) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))) )) \<le> N3"
  using seq_equal_term_nosimp_entire_bounded
  apply(subgoal_tac " \<exists>N3. \<forall>s.  rsize  (RALTS  ((RSEQ (rders_simp r1 s) r2) #
           (rdistinct  (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s)) {}) ) ) \<le> N3")
  apply(erule exE)
   prefer 2
  apply blast
  apply(subgoal_tac "\<forall>s. rsize (rsimp (RALTS  ((RSEQ (rders_simp r1 s) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))) )) \<le> 
rsize  (RALTS  ((RSEQ (rders_simp r1 s) r2) #
           (rdistinct (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s)) {})  ) ) ")
   prefer 2
  using alts_simp_bounded_by_sloppiest_version apply blast
  apply(rule_tac x = "Suc N3 " in exI)
  apply(rule allI)

  apply(subgoal_tac " rsize
             (rsimp
               (RALTS
                 (RSEQ (rders_simp r1 s) r2 # rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s))))
            \<le> rsize
                (RALTS
                  (RSEQ (rders_simp r1 s) r2 #
                   rdistinct (rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s)) {}))")
   prefer 2
   apply presburger
  apply(subgoal_tac "  rsize
                (RALTS
                  (RSEQ (rders_simp r1 s) r2 #
                   rdistinct (rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s)) {})) \<le> N3")
  
   apply linarith
  apply simp
  done



lemma M1seq:
  fixes r1 r2 
  shows "(\<forall>s. rsize (rders_simp r1 s) < N1 \<and> rsize (rders_simp r2 s) < N2)
           \<Longrightarrow> \<exists>N3.\<forall>s.(rsize (rders_simp (RSEQ r1 r2) s)) < N3"
  apply(frule seq_equal_term_entire_bounded)
  apply(erule exE)
  apply(rule_tac x = "max (N3+2) (Suc (Suc (rsize r1) + (rsize r2)))" in exI)
  apply(rule allI)
  apply(case_tac "s = []")
   prefer 2
   apply (metis add_2_eq_Suc' le_imp_less_Suc less_SucI max.strict_coboundedI1 shape_derssimp_seq)
  by (metis add.assoc less_Suc_eq less_max_iff_disj plus_1_eq_Suc rders_simp.simps(1) rsize.simps(5))

 (*  apply (simp add: less_SucI shape_derssimp_seq(2))
   apply (meson less_SucI less_max_iff_disj)
  apply simp
  done*)

(*lemma empty_diff:
  shows "s = [] \<Longrightarrow>
        (rsize (rders_simp (RSEQ r1 r2) s)) \<le> 
        (max 
        (rsize (rsimp (RALTS (RSEQ (rders r1 s) r2 # rders_cond_list r2 (nullable_bools r1 (orderedPref s)) (orderedSuf s)))))
        (Suc (rsize r1 + rsize r2)) ) "
  apply simp
  done*)
(*For star related  bound*)

lemma star_is_a_singleton_list_derc:
  shows " \<exists>Ss.  rders_simp (RSTAR r) [c] = rsimp (RALTS ( (map (\<lambda>s1.  rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)))"
  apply simp
  apply(rule_tac x = "[[c]]" in exI)
  apply auto
  apply(case_tac "rsimp (rder c r)")
       apply simp
      apply auto
   apply(subgoal_tac "rsimp (RSEQ x41 x42) = RSEQ x41 x42")
  prefer 2
    apply (metis rsimp_idem)
   apply(subgoal_tac "rsimp x41 = x41")
  prefer 2
  using rsimp_inner_idem1 apply blast
   apply(subgoal_tac "rsimp x42 = x42")
  prefer 2
  using rsimp_inner_idem1 apply blast
   apply simp
  apply(subgoal_tac "map rsimp x5 = x5")
  prefer 2
  using rsimp_inner_idem3 apply blast
  apply simp
  apply(subgoal_tac "rflts x5 = x5")
   prefer 2
  using rsimp_inner_idem4 apply blast
  apply simp
  using rsimp_inner_idem4 by auto
  


lemma rder_rsimp_ALTs_commute:
  shows "  (rder x (rsimp_ALTs rs)) = rsimp_ALTs (map (rder x) rs)"
  apply(induct rs)
   apply simp
  apply(case_tac rs)
   apply simp
  apply auto
  done



fun star_update :: "char \<Rightarrow> rrexp \<Rightarrow> char list list => char list list" where
"star_update c r [] = []"
|"star_update c r (s # Ss) = (if (rnullable (rders_simp r s)) 
                                then (s@[c]) # [c] # (star_update c r Ss) 
                               else   (s@[c]) # (star_update c r Ss) )"

lemma star_update_case1:
  shows "rnullable (rders_simp r s) \<Longrightarrow> star_update c r (s # Ss) = (s @ [c]) # [c] # (star_update c r Ss)"
  
  by force

lemma star_update_case2:
  shows "\<not>rnullable (rders_simp r s) \<Longrightarrow> star_update c r (s # Ss) = (s @ [c]) # (star_update c r Ss)"
  by simp

lemma bubble_break: shows "rflts [r, RZERO] = rflts [r]"
  apply(case_tac r)
       apply simp+
  done

lemma rsimp_alts_idem_aux1:
  shows "rsimp_ALTs (rdistinct (rflts [rsimp a]) {}) = rsimp (RALTS [a])"
  by force



lemma rsimp_alts_idem_aux2:
  shows "rsimp a = rsimp (RALTS [a])"
  apply(simp)
  apply(case_tac "rsimp a")
       apply simp+
  apply (metis no_alt_short_list_after_simp no_further_dB_after_simp)
  by simp

lemma rsimp_alts_idem:
  shows "rsimp (rsimp_ALTs (a # as)) = rsimp (rsimp_ALTs (a # [(rsimp (rsimp_ALTs as))] ))"
  apply(induct as)
   apply(subgoal_tac "rsimp (rsimp_ALTs [a, rsimp (rsimp_ALTs [])]) = rsimp (rsimp_ALTs [a, RZERO])")
  prefer 2
    apply simp
  using bubble_break rsimp_alts_idem_aux2 apply auto[1]
  apply(case_tac as)
   apply(subgoal_tac "rsimp_ALTs( aa # as) = aa")
  prefer 2
    apply simp
  using head_one_more_simp apply fastforce
  apply(subgoal_tac "rsimp_ALTs (aa # as) = RALTS (aa # as)")
  prefer 2
  
  using rsimp_ALTs.simps(3) apply presburger
  
  apply(simp only:)
  apply(subgoal_tac "rsimp_ALTs (a # aa # aaa # list) = RALTS (a # aa # aaa # list)")
  prefer 2
  using rsimp_ALTs.simps(3) apply presburger
  apply(simp only:)
  apply(subgoal_tac "rsimp_ALTs [a, rsimp (RALTS (aa # aaa # list))] = RALTS (a # [rsimp (RALTS (aa # aaa # list))])")
  prefer 2
  
  using rsimp_ALTs.simps(3) apply presburger
  apply(simp only:)
  using simp_flatten2
  apply(subgoal_tac " rsimp (RALT a (rsimp (RALTS (aa # aaa # list))))  =  rsimp (RALT a ((RALTS (aa # aaa # list)))) ")
  prefer 2

  apply (metis head_one_more_simp list.simps(9) rsimp.simps(2))
  apply (simp only:)
  done


lemma rsimp_alts_idem2:
  shows "rsimp (rsimp_ALTs (a # as)) = rsimp (rsimp_ALTs ((rsimp a) # [(rsimp (rsimp_ALTs as))] ))"
  using head_one_more_simp rsimp_alts_idem by auto


lemma evolution_step1:
  shows "rsimp
        (rsimp_ALTs
          (rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)) # map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) =
         rsimp 
        (rsimp_ALTs
          (rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)) # [(rsimp (rsimp_ALTs (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)))]))   "
  using rsimp_alts_idem by auto

lemma evolution_step2:
  assumes " rsimp (rsimp_ALTs (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) =
       rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))"
  shows "rsimp 
        (rsimp_ALTs 
          (rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)) # map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) = 
                 rsimp 
        (rsimp_ALTs
          (rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)) # [ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]))  "
  by (simp add: assms rsimp_alts_idem)

lemma rsimp_seq_aux1:
  shows "r = RONE \<and> r2 = RSTAR r0 \<Longrightarrow> rsimp_SEQ r r2 = r2"
  apply simp
  done

lemma multiple_alts_simp_flatten:
  shows "rsimp (RALT (RALT r1 r2) (rsimp_ALTs rs)) = rsimp (RALTS (r1 # r2 # rs))"
  by (metis Cons_eq_appendI append_self_conv2 rsimp_ALTs.simps(2) rsimp_ALTs.simps(3) rsimp_alts_idem simp_flatten)


lemma evo3_main_aux1:
  shows "rsimp
            (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
              (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))) =
           rsimp
            (RALTS
              (RSEQ (rders_simp r (a @ [x])) (RSTAR r) #
               RSEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))"
  apply(subgoal_tac "rsimp
            (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
              (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))) =
rsimp
            (RALT (RALT (RSEQ ( (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
              (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))) ")
  prefer 2
   apply (simp add: rsimp_idem)
  apply (simp only:)
  apply(subst multiple_alts_simp_flatten)
  by simp


lemma evo3_main_nullable:
  shows "
\<And>a Ss.
       \<lbrakk>rsimp (rsimp_ALTs (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) =
        rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)));
        rders_simp r a \<noteq> RONE; rders_simp r a \<noteq> RZERO; rnullable (rders_simp r a)\<rbrakk>
       \<Longrightarrow> rsimp
            (rsimp_ALTs
              [rder x (RSEQ (rders_simp r a) (RSTAR r)),
               rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) =
           rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r (a # Ss))))"
  apply(subgoal_tac "rder x (RSEQ (rders_simp r a) (RSTAR r)) 
                   = RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ (rder x r) (RSTAR r))")
  prefer 2
   apply simp
  apply(simp only:)
  apply(subgoal_tac "star_update x r (a # Ss) = (a @ [x]) # [x] # (star_update x r Ss)")
   prefer 2
  using star_update_case1 apply presburger
  apply(simp only:)
  apply(subst List.list.map(2))+
  apply(subgoal_tac "rsimp
            (rsimp_ALTs
              [RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ (rder x r) (RSTAR r)),
               rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) = 
rsimp
            (RALTS
              [RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ (rder x r) (RSTAR r)),
               rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))])")
  prefer 2
  using rsimp_ALTs.simps(3) apply presburger
  apply(simp only:)
  apply(subgoal_tac " rsimp
            (rsimp_ALTs
              (rsimp_SEQ (rders_simp r (a @ [x])) (RSTAR r) #
               rsimp_SEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))) 
= 
 rsimp
            (RALTS
              (rsimp_SEQ (rders_simp r (a @ [x])) (RSTAR r) #
               rsimp_SEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))")

  prefer 2
  using rsimp_ALTs.simps(3) apply presburger
  apply (simp only:)
  apply(subgoal_tac " rsimp
            (RALT (RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ ( (rder x r)) (RSTAR r)))
              (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) = 
             rsimp
            (RALT (RALT (RSEQ (rsimp (rder x (rders_simp r a))) (RSTAR r)) (RSEQ (rsimp (rder x r)) (RSTAR r)))
              (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))))")
  prefer 2
   apply (simp add: rsimp_idem)
  apply(simp only:)
  apply(subgoal_tac "             rsimp
            (RALT (RALT (RSEQ (rsimp (rder x (rders_simp r a))) (RSTAR r)) (RSEQ (rsimp (rder x r)) (RSTAR r)))
              (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) = 
             rsimp
            (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
              (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))))")
   prefer 2
  using rders_simp_append rders_simp_one_char rsimp_idem apply presburger
  apply(simp only:)
  apply(subgoal_tac " rsimp
            (RALTS
              (rsimp_SEQ (rders_simp r (a @ [x])) (RSTAR r) #
               rsimp_SEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))) = 
 rsimp
            (RALTS
              (RSEQ (rders_simp r (a @ [x])) (RSTAR r) #
               RSEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))")
  prefer 2
  apply (smt (z3) idiot2 list.simps(9) rrexp.distinct(9) rsimp.simps(1) rsimp.simps(2) rsimp.simps(3) rsimp.simps(4) rsimp.simps(6) rsimp_idem)
  apply(simp only:)
  apply(subgoal_tac "      rsimp
            (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
              (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) =
     rsimp
            (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
              ( (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))))  ")
  prefer 2
  using rsimp_idem apply force
  apply(simp only:)
  using evo3_main_aux1 by blast
  

lemma evo3_main_not1:
  shows " \<not>rnullable (rders_simp r a) \<Longrightarrow> rder x (RSEQ (rders_simp r a) (RSTAR r)) = RSEQ (rder x (rders_simp r a)) (RSTAR r)"
  by fastforce


lemma evo3_main_not2:
  shows "\<not>rnullable (rders_simp r a) \<Longrightarrow>  rsimp
            (rsimp_ALTs
              (rder x (RSEQ (rders_simp r a) (RSTAR r)) # rs)) = rsimp
            (rsimp_ALTs
              ((RSEQ (rders_simp r (a @ [x])) (RSTAR r)) # rs))"
  by (simp add: rders_simp_append rsimp_alts_idem2 rsimp_idem)

lemma evo3_main_not3:
  shows "rsimp
            (rsimp_ALTs
              (rsimp_SEQ r1 (RSTAR r) # rs)) = 
         rsimp (rsimp_ALTs
              (RSEQ r1 (RSTAR r) # rs))"
  by (metis idiot2 rrexp.distinct(9) rsimp.simps(1) rsimp.simps(3) rsimp.simps(4) rsimp.simps(6) rsimp_alts_idem rsimp_alts_idem2)


lemma evo3_main_notnullable:
  shows "\<And>a Ss.
       \<lbrakk>rsimp (rsimp_ALTs (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) =
        rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)));
        rders_simp r a \<noteq> RONE; rders_simp r a \<noteq> RZERO; \<not>rnullable (rders_simp r a)\<rbrakk>
       \<Longrightarrow> rsimp
            (rsimp_ALTs
              [rder x (RSEQ (rders_simp r a) (RSTAR r)),
               rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) =
           rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r (a # Ss))))"
  apply(subst star_update_case2)
   apply simp
  apply(subst List.list.map(2))
  apply(subst evo3_main_not2)
   apply simp
  apply(subst evo3_main_not3)
  using rsimp_alts_idem by presburger


lemma evo3_aux2:
  shows "rders_simp r a = RONE \<Longrightarrow> rsimp_SEQ (rders_simp (rders_simp r a) [x]) (RSTAR r) = RZERO"
  by simp
lemma evo3_aux3:
  shows "rsimp (rsimp_ALTs (RZERO # rs)) = rsimp (rsimp_ALTs rs)"
  by (metis list.simps(8) list.simps(9) rdistinct.simps(1) rflts.simps(1) rflts.simps(2) rsimp.simps(2) rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3) rsimp_alts_idem)

lemma evo3_aux4:
  shows " rsimp
            (rsimp_ALTs
              [RSEQ (rder x r) (RSTAR r),
               rsimp (rsimp_ALTs rs)]) =
           rsimp
            (rsimp_ALTs
              (rsimp_SEQ (rders_simp r [x]) (RSTAR r) # rs))"
  by (metis rders_simp_one_char rsimp.simps(1) rsimp.simps(6) rsimp_alts_idem rsimp_alts_idem2)

lemma evo3_aux5:
  shows "rders_simp r a \<noteq> RONE \<and> rders_simp r a \<noteq> RZERO \<Longrightarrow> rsimp_SEQ (rders_simp r a) (RSTAR r) = RSEQ (rders_simp r a) (RSTAR r)"
  using idiot2 by blast


lemma evolution_step3:
  shows" \<And>a Ss.
       rsimp (rsimp_ALTs (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) =
       rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))) \<Longrightarrow>
       rsimp
        (rsimp_ALTs
          [rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)),
           rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) =
       rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r (a # Ss))))"
  apply(case_tac "rders_simp r a = RONE")
   apply(subst rsimp_seq_aux1)
    apply simp
  apply(subst rder.simps(6))
   apply(subgoal_tac "rnullable (rders_simp r a)")
    prefer 2
  using rnullable.simps(2) apply presburger
   apply(subst star_update_case1)
    apply simp

   apply(subst List.list.map)+
  apply(subst rders_simp_append)
   apply(subst evo3_aux2)
    apply simp
   apply(subst evo3_aux3)
   apply(subst evo3_aux4)
   apply simp
  apply(case_tac "rders_simp r a = RZERO")

   apply (simp add: rsimp_alts_idem2)
   apply(subgoal_tac "rders_simp r (a @ [x]) = RZERO")
  prefer 2
  using rder.simps(1) rders_simp_append rders_simp_one_char rsimp.simps(3) apply presburger
  using rflts.simps(2) rsimp.simps(3) rsimp_SEQ.simps(1) apply presburger
  apply(subst evo3_aux5)
   apply simp
  apply(case_tac "rnullable (rders_simp r a) ")
  using evo3_main_nullable apply blast
  using evo3_main_notnullable apply blast
  done

(*
proof (prove)
goal (1 subgoal):
 1. map f (a # s) = f a # map f s 
Auto solve_direct: the current goal can be solved directly with
  HOL.nitpick_simp(115): map ?f (?x21.0 # ?x22.0) = ?f ?x21.0 # map ?f ?x22.0
  List.list.map(2): map ?f (?x21.0 # ?x22.0) = ?f ?x21.0 # map ?f ?x22.0
  List.list.simps(9): map ?f (?x21.0 # ?x22.0) = ?f ?x21.0 # map ?f ?x22.0
*)
lemma starseq_list_evolution:
  fixes  r :: rrexp and Ss :: "char list list" and x :: char 
  shows "rsimp (rsimp_ALTs (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss) ) =
         rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))  )"   
  apply(induct Ss)
   apply simp
  apply(subst List.list.map(2))
  apply(subst evolution_step2)
   apply simp


  sorry


lemma star_seqs_produce_star_seqs:
  shows "rsimp (rsimp_ALTs (map (rder x \<circ> (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss))
       = rsimp (rsimp_ALTs (map ( (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r)))) Ss))"
  by (meson comp_apply)

lemma map_der_lambda_composition:
  shows "map (rder x) (map (\<lambda>s. f s) Ss) = map (\<lambda>s. (rder x (f s))) Ss"
  by force

lemma ralts_vs_rsimpalts:
  shows "rsimp (RALTS rs) = rsimp (rsimp_ALTs rs)"
  by (metis evo3_aux3 rsimp_ALTs.simps(2) rsimp_ALTs.simps(3) simp_flatten2)
  

lemma linearity_of_list_of_star_or_starseqs: 
  fixes r::rrexp and Ss::"char list list" and x::char
  shows "\<exists>Ssa. rsimp (rder x (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))) =
                 rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ssa)))"
  apply(subst rder_rsimp_ALTs_commute)
  apply(subst map_der_lambda_composition)
  using starseq_list_evolution
  apply(rule_tac x = "star_update x r Ss" in exI)
  apply(subst ralts_vs_rsimpalts)
  by simp



(*certified correctness---does not depend on any previous sorry*)
lemma star_list_push_der: shows  " \<lbrakk>xs \<noteq> [] \<Longrightarrow> \<exists>Ss. rders_simp (RSTAR r) xs = rsimp (RALTS (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss));
        xs @ [x] \<noteq> []; xs \<noteq> []\<rbrakk> \<Longrightarrow>
     \<exists>Ss. rders_simp (RSTAR r ) (xs @ [x]) = 
        rsimp (RALTS (map (\<lambda>s1. (rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) ) Ss) )"
  apply(subgoal_tac  "\<exists>Ss. rders_simp (RSTAR r) xs = rsimp (RALTS (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))")
  prefer 2
  apply blast
  apply(erule exE)
  apply(subgoal_tac "rders_simp (RSTAR r) (xs @ [x]) = rsimp (rder x (rsimp (RALTS (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))))")
  prefer 2
  using rders_simp_append
  using rders_simp_one_char apply presburger
  apply(rule_tac x= "Ss" in exI)
  apply(subgoal_tac " rsimp (rder x (rsimp (RALTS (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)))) = 
                       rsimp (rsimp (rder x (RALTS (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))))")
  prefer 2
  using inside_simp_removal rsimp_idem apply presburger
  apply(subgoal_tac "rsimp (rsimp (rder x (RALTS (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)))) =
                     rsimp (rsimp (RALTS (map (rder x) (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))))")
  prefer 2
  using rder.simps(4) apply presburger
  apply(subgoal_tac "rsimp (rsimp (RALTS (map (rder x) (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)))) =
                     rsimp (rsimp (RALTS (map (\<lambda>s1. (rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r)))) Ss)))")
   apply (metis rsimp_idem)
  by (metis map_der_lambda_composition)

lemma simp_in_lambdas : 
  shows "
rsimp (RALTS (map (\<lambda>s1. (rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) ) Ss) ) = 
rsimp (RALTS (map (\<lambda>s1. (rsimp (rder x  (rsimp_SEQ (rders_simp r s1) (RSTAR r))))) Ss))"
  by (metis (no_types, lifting) comp_apply list.map_comp map_eq_conv rsimp.simps(2) rsimp_idem)


lemma starder_is_a_list_of_stars_or_starseqs:
  shows "s \<noteq> [] \<Longrightarrow> \<exists>Ss.  rders_simp (RSTAR r) s = rsimp (RALTS( (map (\<lambda>s1.  rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)))"
  apply(induct s rule: rev_induct)
  apply simp
  apply(case_tac "xs = []")
  using star_is_a_singleton_list_derc
   apply(simp)
  apply(subgoal_tac "\<exists>Ss. rders_simp (RSTAR r) (xs @ [x]) = 
          rsimp (RALTS (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss))")
  prefer 2
  using star_list_push_der apply presburger


  by (metis ralts_vs_rsimpalts starseq_list_evolution)


lemma starder_is_a_list:
  shows " \<exists>Ss. rders_simp (RSTAR r) s = rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))) \<or> rders_simp (RSTAR r) s = RSTAR r"
  apply(case_tac s)
  prefer 2
  apply (metis neq_Nil_conv starder_is_a_list_of_stars_or_starseqs)
  apply simp
  done


(** start about bounds here**)


lemma list_simp_size:
  shows "rlist_size (map rsimp rs) \<le> rlist_size rs"
  apply(induct rs)
   apply simp
  apply simp
  apply (subgoal_tac "rsize (rsimp a) \<le> rsize a")
  prefer 2
  using rsimp_mono apply fastforce
  using add_le_mono by presburger

lemma inside_list_simp_inside_list:
  shows "r \<in> set rs \<Longrightarrow> rsimp r \<in> set (map rsimp rs)"
  apply (induct rs)
   apply simp
  apply auto
  done


lemma rsize_star_seq_list:
  shows "(\<forall>s. rsize (rders_simp r0 s) < N0 ) \<Longrightarrow>  \<exists>N3.\<forall>Ss.  
rlist_size (rdistinct (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss) {}) < N3"
  sorry


lemma rdistinct_bound_by_no_simp:
  shows "

   rlist_size (rdistinct  (map rsimp rs) (set (map rsimp ss)))
 \<le> (rlist_size (rdistinct rs (set ss)))
"
  apply(induct rs arbitrary: ss)
   apply simp
  apply(case_tac "a \<in> set ss")
   apply(subgoal_tac "rsimp a \<in> set (map rsimp ss)")
  prefer 2
  using inside_list_simp_inside_list apply blast

   apply simp
  apply simp
  by (metis List.set_insert add_le_mono image_insert insert_absorb rsimp_mono trans_le_add2)


lemma starder_closed_form_bound_aux1:
  shows 
"\<forall>Ss. rsize (rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) \<le>
  Suc (rlist_size ( (rdistinct ( ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss))) {}))) "
  
  sorry

lemma starder_closed_form_bound:
  shows "(\<forall>s. rsize (rders_simp r0 s) < N0 ) \<Longrightarrow> \<exists>N3.\<forall>Ss.  
rsize(rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) < N3"
  apply(subgoal_tac " \<exists>N3.\<forall>Ss.  
rlist_size (rdistinct (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss) {}) < N3")
  prefer 2
  
  using rsize_star_seq_list apply auto[1]
  apply(erule exE)
  apply(rule_tac x = "Suc N3" in exI)
  apply(subgoal_tac "\<forall>Ss. rsize (rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) \<le>
 Suc (rlist_size ( (rdistinct ( ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss))) {})))")
  prefer 2
  using starder_closed_form_bound_aux1 apply blast
  by (meson less_trans_Suc linorder_not_le not_less_eq)

  
thm starder_closed_form_bound_aux1

(*
 "ralts_vs_rsimpalts", , and "starder_closed_form_bound_aux1", which could be due to a bug in Sledgehammer or to inconsistent axioms (including "sorry"s) 
*)

lemma starder_size_bound:
  shows "(\<forall>s. rsize (rders_simp r0 s) < N0 ) \<Longrightarrow> \<exists>N3.\<forall>Ss.  
rsize(rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) < N3 \<and>
rsize (RSTAR r0) < N3"
  apply(subgoal_tac " \<exists>N3.\<forall>Ss.  
rsize(rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) < N3")
  prefer 2
  using starder_closed_form_bound apply blast
  apply(erule exE)
  apply(rule_tac x = "max N3 (Suc (rsize (RSTAR r0)))" in exI)
  using less_max_iff_disj by blast




lemma finite_star:
  shows "(\<forall>s. rsize (rders_simp r0 s) < N0 )
           \<Longrightarrow> \<exists>N3. \<forall>s.(rsize (rders_simp (RSTAR r0) s)) < N3"
  apply(subgoal_tac  " \<exists>N3. \<forall>Ss. 
rsize(rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) < N3 \<and>
rsize (RSTAR r0) < N3")
  prefer 2
  using starder_size_bound apply blast
  apply(erule exE)
  apply(rule_tac x = N3 in exI)
  by (metis starder_is_a_list)


lemma rderssimp_zero:
  shows"rders_simp RZERO s = RZERO"
  apply(induction s)
  apply simp
  by simp

lemma rderssimp_one:
  shows"rders_simp RONE (a # s) = RZERO"
  apply(induction s)
  apply simp
  by simp

lemma rderssimp_char:
  shows "rders_simp (RCHAR c) s = RONE \<or> rders_simp (RCHAR c) s = RZERO \<or> rders_simp (RCHAR c) s = (RCHAR c)"
  apply auto
  by (metis rder.simps(2) rder.simps(3) rders_simp.elims rders_simp.simps(2) rderssimp_one rsimp.simps(4))

lemma finite_size_ders:
  fixes r
  shows " \<exists>Nr. \<forall>s. rsize (rders_simp r s) < Nr"
  apply(induct r rule: rrexp.induct)
       apply auto
  apply(rule_tac x = "2" in exI)
  using rderssimp_zero rsize.simps(1) apply presburger
      apply(rule_tac x = "2" in exI)
      apply (metis Suc_1 lessI rders_simp.elims rderssimp_one rsize.simps(1) rsize.simps(2))
     apply(rule_tac x = "2" in meta_spec)
     apply (metis lessI rderssimp_char rsize.simps(1) rsize.simps(2) rsize.simps(3))
  
  using M1seq apply blast
   prefer 2

   apply (simp add: finite_star)
  sorry

lemma finite_list_of_ders:
  fixes r
  shows"\<exists>dersset. ( (finite dersset) \<and> (\<forall>s. (rders_simp r s) \<in> dersset) )"
  sorry



unused_thms
lemma seq_ders_shape:
  shows "E"

  oops

(*rsimp (rders (RSEQ r1 r2) s) =
         rsimp RALT [RSEQ (rders r1 s) r2, rders r2 si, ...]
         where si is the i-th shortest suffix of s such that si \<in> L r2"
*)


end