thys2/SizeBound5CT.thy
author Chengsong
Fri, 11 Feb 2022 17:28:33 +0000
changeset 430 579caa608a15
parent 428 5dcecc92608e
permissions -rw-r--r--
i


theory SizeBound5CT
  imports "Lexer" "PDerivs" 
begin

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

datatype bit = Z | S

fun code :: "val \<Rightarrow> bit list"
where
  "code Void = []"
| "code (Char c) = []"
| "code (Left v) = Z # (code v)"
| "code (Right v) = S # (code v)"
| "code (Seq v1 v2) = (code v1) @ (code v2)"
| "code (Stars []) = [S]"
| "code (Stars (v # vs)) =  (Z # code v) @ code (Stars vs)"


fun 
  Stars_add :: "val \<Rightarrow> val \<Rightarrow> val"
where
  "Stars_add v (Stars vs) = Stars (v # vs)"

function
  decode' :: "bit list \<Rightarrow> rexp \<Rightarrow> (val * bit list)"
where
  "decode' bs ZERO = (undefined, bs)"
| "decode' bs ONE = (Void, bs)"
| "decode' bs (CH d) = (Char d, bs)"
| "decode' [] (ALT r1 r2) = (Void, [])"
| "decode' (Z # bs) (ALT r1 r2) = (let (v, bs') = decode' bs r1 in (Left v, bs'))"
| "decode' (S # bs) (ALT r1 r2) = (let (v, bs') = decode' bs r2 in (Right v, bs'))"
| "decode' bs (SEQ r1 r2) = (let (v1, bs') = decode' bs r1 in
                             let (v2, bs'') = decode' bs' r2 in (Seq v1 v2, bs''))"
| "decode' [] (STAR r) = (Void, [])"
| "decode' (S # bs) (STAR r) = (Stars [], bs)"
| "decode' (Z # bs) (STAR r) = (let (v, bs') = decode' bs r in
                                    let (vs, bs'') = decode' bs' (STAR r) 
                                    in (Stars_add v vs, bs''))"
by pat_completeness auto

lemma decode'_smaller:
  assumes "decode'_dom (bs, r)"
  shows "length (snd (decode' bs r)) \<le> length bs"
using assms
apply(induct bs r)
apply(auto simp add: decode'.psimps split: prod.split)
using dual_order.trans apply blast
by (meson dual_order.trans le_SucI)

termination "decode'"  
apply(relation "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))") 
apply(auto dest!: decode'_smaller)
by (metis less_Suc_eq_le snd_conv)

definition
  decode :: "bit list \<Rightarrow> rexp \<Rightarrow> val option"
where
  "decode ds r \<equiv> (let (v, ds') = decode' ds r 
                  in (if ds' = [] then Some v else None))"

lemma decode'_code_Stars:
  assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> (\<forall>x. decode' (code v @ x) r = (v, x)) \<and> flat v \<noteq> []" 
  shows "decode' (code (Stars vs) @ ds) (STAR r) = (Stars vs, ds)"
  using assms
  apply(induct vs)
  apply(auto)
  done

lemma decode'_code:
  assumes "\<Turnstile> v : r"
  shows "decode' ((code v) @ ds) r = (v, ds)"
using assms
  apply(induct v r arbitrary: ds) 
  apply(auto)
  using decode'_code_Stars by blast

lemma decode_code:
  assumes "\<Turnstile> v : r"
  shows "decode (code v) r = Some v"
  using assms unfolding decode_def
  by (smt append_Nil2 decode'_code old.prod.case)


section {* Annotated Regular Expressions *}

datatype arexp = 
  AZERO
| AONE "bit list"
| ACHAR "bit list" char
| ASEQ "bit list" arexp arexp
| AALTs "bit list" "arexp list"
| ASTAR "bit list" arexp

abbreviation
  "AALT bs r1 r2 \<equiv> AALTs bs [r1, r2]"

fun asize :: "arexp \<Rightarrow> nat" where
  "asize AZERO = 1"
| "asize (AONE cs) = 1" 
| "asize (ACHAR cs c) = 1"
| "asize (AALTs cs rs) = Suc (sum_list (map asize rs))"
| "asize (ASEQ cs r1 r2) = Suc (asize r1 + asize r2)"
| "asize (ASTAR cs r) = Suc (asize r)"

fun 
  erase :: "arexp \<Rightarrow> rexp"
where
  "erase AZERO = ZERO"
| "erase (AONE _) = ONE"
| "erase (ACHAR _ c) = CH c"
| "erase (AALTs _ []) = ZERO"
| "erase (AALTs _ [r]) = (erase r)"
| "erase (AALTs bs (r#rs)) = ALT (erase r) (erase (AALTs bs rs))"
| "erase (ASEQ _ r1 r2) = SEQ (erase r1) (erase r2)"
| "erase (ASTAR _ r) = STAR (erase r)"


fun fuse :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp" where
  "fuse bs AZERO = AZERO"
| "fuse bs (AONE cs) = AONE (bs @ cs)" 
| "fuse bs (ACHAR cs c) = ACHAR (bs @ cs) c"
| "fuse bs (AALTs cs rs) = AALTs (bs @ cs) rs"
| "fuse bs (ASEQ cs r1 r2) = ASEQ (bs @ cs) r1 r2"
| "fuse bs (ASTAR cs r) = ASTAR (bs @ cs) r"

lemma fuse_append:
  shows "fuse (bs1 @ bs2) r = fuse bs1 (fuse bs2 r)"
  apply(induct r)
  apply(auto)
  done


fun intern :: "rexp \<Rightarrow> arexp" where
  "intern ZERO = AZERO"
| "intern ONE = AONE []"
| "intern (CH c) = ACHAR [] c"
| "intern (ALT r1 r2) = AALT [] (fuse [Z] (intern r1)) 
                                (fuse [S]  (intern r2))"
| "intern (SEQ r1 r2) = ASEQ [] (intern r1) (intern r2)"
| "intern (STAR r) = ASTAR [] (intern r)"


fun retrieve :: "arexp \<Rightarrow> val \<Rightarrow> bit list" where
  "retrieve (AONE bs) Void = bs"
| "retrieve (ACHAR bs c) (Char d) = bs"
| "retrieve (AALTs bs [r]) v = bs @ retrieve r v"
| "retrieve (AALTs bs (r#rs)) (Left v) = bs @ retrieve r v"
| "retrieve (AALTs bs (r#rs)) (Right v) = bs @ retrieve (AALTs [] rs) v"
| "retrieve (ASEQ bs r1 r2) (Seq v1 v2) = bs @ retrieve r1 v1 @ retrieve r2 v2"
| "retrieve (ASTAR bs r) (Stars []) = bs @ [S]"
| "retrieve (ASTAR bs r) (Stars (v#vs)) = 
     bs @ [Z] @ retrieve r v @ retrieve (ASTAR [] r) (Stars vs)"



fun
 bnullable :: "arexp \<Rightarrow> bool"
where
  "bnullable (AZERO) = False"
| "bnullable (AONE bs) = True"
| "bnullable (ACHAR bs c) = False"
| "bnullable (AALTs bs rs) = (\<exists>r \<in> set rs. bnullable r)"
| "bnullable (ASEQ bs r1 r2) = (bnullable r1 \<and> bnullable r2)"
| "bnullable (ASTAR bs r) = True"

abbreviation
  bnullables :: "arexp list \<Rightarrow> bool"
where
  "bnullables rs \<equiv> (\<exists>r \<in> set rs. bnullable r)"

fun 
  bmkeps :: "arexp \<Rightarrow> bit list" and
  bmkepss :: "arexp list \<Rightarrow> bit list"
where
  "bmkeps(AONE bs) = bs"
| "bmkeps(ASEQ bs r1 r2) = bs @ (bmkeps r1) @ (bmkeps r2)"
| "bmkeps(AALTs bs rs) = bs @ (bmkepss rs)"
| "bmkeps(ASTAR bs r) = bs @ [S]"
| "bmkepss [] = []"
| "bmkepss (r # rs) = (if bnullable(r) then (bmkeps r) else (bmkepss rs))"

lemma bmkepss1:
  assumes "\<not> bnullables rs1"
  shows "bmkepss (rs1 @ rs2) = bmkepss rs2"
  using assms
  by (induct rs1) (auto)

lemma bmkepss2:
  assumes "bnullables rs1"
  shows "bmkepss (rs1 @ rs2) = bmkepss rs1"
  using assms
  by (induct rs1) (auto)


fun
 bder :: "char \<Rightarrow> arexp \<Rightarrow> arexp"
where
  "bder c (AZERO) = AZERO"
| "bder c (AONE bs) = AZERO"
| "bder c (ACHAR bs d) = (if c = d then AONE bs else AZERO)"
| "bder c (AALTs bs rs) = AALTs bs (map (bder c) rs)"
| "bder c (ASEQ bs r1 r2) = 
     (if bnullable r1
      then AALT bs (ASEQ [] (bder c r1) r2) (fuse (bmkeps r1) (bder c r2))
      else ASEQ bs (bder c r1) r2)"
| "bder c (ASTAR bs r) = ASEQ bs (fuse [Z] (bder c r)) (ASTAR [] r)"


fun 
  bders :: "arexp \<Rightarrow> string \<Rightarrow> arexp"
where
  "bders r [] = r"
| "bders r (c#s) = bders (bder c r) s"

lemma bders_append:
  "bders c (s1 @ s2) = bders (bders c s1) s2"
  apply(induct s1 arbitrary: c s2)
  apply(simp_all)
  done

lemma bnullable_correctness:
  shows "nullable (erase r) = bnullable r"
  apply(induct r rule: erase.induct)
  apply(simp_all)
  done

lemma erase_fuse:
  shows "erase (fuse bs r) = erase r"
  apply(induct r rule: erase.induct)
  apply(simp_all)
  done

lemma erase_intern [simp]:
  shows "erase (intern r) = r"
  apply(induct r)
  apply(simp_all add: erase_fuse)
  done

lemma erase_bder [simp]:
  shows "erase (bder a r) = der a (erase r)"
  apply(induct r rule: erase.induct)
  apply(simp_all add: erase_fuse bnullable_correctness)
  done

lemma erase_bders [simp]:
  shows "erase (bders r s) = ders s (erase r)"
  apply(induct s arbitrary: r )
  apply(simp_all)
  done

lemma bnullable_fuse:
  shows "bnullable (fuse bs r) = bnullable r"
  apply(induct r arbitrary: bs)
  apply(auto)
  done

lemma retrieve_encode_STARS:
  assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> code v = retrieve (intern r) v"
  shows "code (Stars vs) = retrieve (ASTAR [] (intern r)) (Stars vs)"
  using assms
  apply(induct vs)
  apply(simp_all)
  done

lemma retrieve_fuse2:
  assumes "\<Turnstile> v : (erase r)"
  shows "retrieve (fuse bs r) v = bs @ retrieve r v"
  using assms
  apply(induct r arbitrary: v bs)
  apply(auto elim: Prf_elims)[4]
  apply(case_tac x2a)
  apply(simp)
  using Prf_elims(1) apply blast
  apply(case_tac x2a)
  apply(simp)
  apply(simp)
  apply(case_tac list)
  apply(simp)
  apply(simp)
  apply (smt (verit, best) Prf_elims(3) append_assoc retrieve.simps(4) retrieve.simps(5))
  apply(simp)
  using retrieve_encode_STARS
  apply(auto elim!: Prf_elims)[1]
  apply(case_tac vs)
  apply(simp)
  apply(simp)
  done

lemma retrieve_fuse:
  assumes "\<Turnstile> v : r"
  shows "retrieve (fuse bs (intern r)) v = bs @ retrieve (intern r) v"
  using assms 
  by (simp_all add: retrieve_fuse2)


lemma retrieve_code:
  assumes "\<Turnstile> v : r"
  shows "code v = retrieve (intern r) v"
  using assms
  apply(induct v r )
  apply(simp_all add: retrieve_fuse retrieve_encode_STARS)
  done


lemma retrieve_AALTs_bnullable1:
  assumes "bnullable r"
  shows "retrieve (AALTs bs (r # rs)) (mkeps (erase (AALTs bs (r # rs))))
         = bs @ retrieve r (mkeps (erase r))"
  using assms
  apply(case_tac rs)
  apply(auto simp add: bnullable_correctness)
  done

lemma retrieve_AALTs_bnullable2:
  assumes "\<not>bnullable r" "bnullables rs"
  shows "retrieve (AALTs bs (r # rs)) (mkeps (erase (AALTs bs (r # rs))))
         = retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs)))"
  using assms
  apply(induct rs arbitrary: r bs)
  apply(auto)
  using bnullable_correctness apply blast
  apply(case_tac rs)
  apply(auto)
  using bnullable_correctness apply blast
  apply(case_tac rs)
  apply(auto)
  done

lemma bmkeps_retrieve_AALTs: 
  assumes "\<forall>r \<in> set rs. bnullable r \<longrightarrow> bmkeps r = retrieve r (mkeps (erase r))" 
          "bnullables rs"
  shows "bs @ bmkepss rs = retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs)))"
 using assms
  apply(induct rs arbitrary: bs)
  apply(auto)
  using retrieve_AALTs_bnullable1 apply presburger
  apply (metis retrieve_AALTs_bnullable2)
  apply (simp add: retrieve_AALTs_bnullable1)
  by (metis retrieve_AALTs_bnullable2)

    
lemma bmkeps_retrieve:
  assumes "bnullable r"
  shows "bmkeps r = retrieve r (mkeps (erase r))"
  using assms
  apply(induct r)
  apply(auto)  
  using bmkeps_retrieve_AALTs by auto

lemma bder_retrieve:
  assumes "\<Turnstile> v : der c (erase r)"
  shows "retrieve (bder c r) v = retrieve r (injval (erase r) c v)"
  using assms  
  apply(induct r arbitrary: v rule: erase.induct)
  using Prf_elims(1) apply auto[1]
  using Prf_elims(1) apply auto[1]
  apply(auto)[1]
  apply (metis Prf_elims(4) injval.simps(1) retrieve.simps(1) retrieve.simps(2))
  using Prf_elims(1) apply blast
  (* AALTs case *)
  apply(simp)
  apply(erule Prf_elims)
  apply(simp)
  apply(simp)
  apply(rename_tac "r\<^sub>1" "r\<^sub>2" rs v)
  apply(erule Prf_elims)
  apply(simp)
  apply(simp)
  apply(case_tac rs)
  apply(simp)
  apply(simp)
  using Prf_elims(3) apply fastforce
  (* ASEQ case *) 
  apply(simp)
  apply(case_tac "nullable (erase r1)")
  apply(simp)
  apply(erule Prf_elims)
  using Prf_elims(2) bnullable_correctness apply force
  apply (simp add: bmkeps_retrieve bnullable_correctness retrieve_fuse2)
  apply (simp add: bmkeps_retrieve bnullable_correctness retrieve_fuse2)
  using Prf_elims(2) apply force
  (* ASTAR case *)  
  apply(rename_tac bs r v)
  apply(simp)  
  apply(erule Prf_elims)
  apply(clarify)
  apply(erule Prf_elims)
  apply(clarify)
  by (simp add: retrieve_fuse2)


lemma MAIN_decode:
  assumes "\<Turnstile> v : ders s r"
  shows "Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r"
  using assms
proof (induct s arbitrary: v rule: rev_induct)
  case Nil
  have "\<Turnstile> v : ders [] r" by fact
  then have "\<Turnstile> v : r" by simp
  then have "Some v = decode (retrieve (intern r) v) r"
    using decode_code retrieve_code by auto
  then show "Some (flex r id [] v) = decode (retrieve (bders (intern r) []) v) r"
    by simp
next
  case (snoc c s v)
  have IH: "\<And>v. \<Turnstile> v : ders s r \<Longrightarrow> 
     Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r" by fact
  have asm: "\<Turnstile> v : ders (s @ [c]) r" by fact
  then have asm2: "\<Turnstile> injval (ders s r) c v : ders s r" 
    by (simp add: Prf_injval ders_append)
  have "Some (flex r id (s @ [c]) v) = Some (flex r id s (injval (ders s r) c v))"
    by (simp add: flex_append)
  also have "... = decode (retrieve (bders (intern r) s) (injval (ders s r) c v)) r"
    using asm2 IH by simp
  also have "... = decode (retrieve (bder c (bders (intern r) s)) v) r"
    using asm by (simp_all add: bder_retrieve ders_append)
  finally show "Some (flex r id (s @ [c]) v) = 
                 decode (retrieve (bders (intern r) (s @ [c])) v) r" by (simp add: bders_append)
qed

definition blexer where
 "blexer r s \<equiv> if bnullable (bders (intern r) s) then 
                decode (bmkeps (bders (intern r) s)) r else None"

lemma blexer_correctness:
  shows "blexer r s = lexer r s"
proof -
  { define bds where "bds \<equiv> bders (intern r) s"
    define ds  where "ds \<equiv> ders s r"
    assume asm: "nullable ds"
    have era: "erase bds = ds" 
      unfolding ds_def bds_def by simp
    have mke: "\<Turnstile> mkeps ds : ds"
      using asm by (simp add: mkeps_nullable)
    have "decode (bmkeps bds) r = decode (retrieve bds (mkeps ds)) r"
      using bmkeps_retrieve
      using asm era
      using bnullable_correctness by force 
    also have "... =  Some (flex r id s (mkeps ds))"
      using mke by (simp_all add: MAIN_decode ds_def bds_def)
    finally have "decode (bmkeps bds) r = Some (flex r id s (mkeps ds))" 
      unfolding bds_def ds_def .
  }
  then show "blexer r s = lexer r s"
    unfolding blexer_def lexer_flex
    by (auto simp add: bnullable_correctness[symmetric])
qed


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

  

fun flts :: "arexp list \<Rightarrow> arexp list"
  where 
  "flts [] = []"
| "flts (AZERO # rs) = flts rs"
| "flts ((AALTs bs  rs1) # rs) = (map (fuse bs) rs1) @ flts rs"
| "flts (r1 # rs) = r1 # flts rs"



fun bsimp_ASEQ :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp \<Rightarrow> arexp"
  where
  "bsimp_ASEQ _ AZERO _ = AZERO"
| "bsimp_ASEQ _ _ AZERO = AZERO"
| "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2"
| "bsimp_ASEQ bs1 r1 r2 = ASEQ  bs1 r1 r2"

lemma bsimp_ASEQ0[simp]:
  shows "bsimp_ASEQ bs r1 AZERO = AZERO"
  by (case_tac r1)(simp_all)

lemma bsimp_ASEQ1:
  assumes "r1 \<noteq> AZERO" "r2 \<noteq> AZERO" "\<nexists>bs. r1 = AONE bs"
  shows "bsimp_ASEQ bs r1 r2 = ASEQ bs r1 r2"
  using assms
  apply(induct bs r1 r2 rule: bsimp_ASEQ.induct)
  apply(auto)
  done

lemma bsimp_ASEQ2[simp]:
  shows "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2"
  by (case_tac r2) (simp_all)


fun bsimp_AALTs :: "bit list \<Rightarrow> arexp list \<Rightarrow> arexp"
  where
  "bsimp_AALTs _ [] = AZERO"
| "bsimp_AALTs bs1 [r] = fuse bs1 r"
| "bsimp_AALTs bs1 rs = AALTs bs1 rs"


fun bsimp :: "arexp \<Rightarrow> arexp" 
  where
  "bsimp (ASEQ bs1 r1 r2) = bsimp_ASEQ bs1 (bsimp r1) (bsimp r2)"
| "bsimp (AALTs bs1 rs) = bsimp_AALTs bs1 (distinctBy (flts (map bsimp rs)) erase {}) "
| "bsimp r = r"


fun 
  bders_simp :: "arexp \<Rightarrow> string \<Rightarrow> arexp"
where
  "bders_simp r [] = r"
| "bders_simp r (c # s) = bders_simp (bsimp (bder c r)) s"

definition blexer_simp where
 "blexer_simp r s \<equiv> if bnullable (bders_simp (intern r) s) then 
                    decode (bmkeps (bders_simp (intern r) s)) r else None"



lemma bders_simp_append:
  shows "bders_simp r (s1 @ s2) = bders_simp (bders_simp r s1) s2"
  apply(induct s1 arbitrary: r s2)
  apply(simp_all)
  done


lemma bmkeps_fuse:
  assumes "bnullable r"
  shows "bmkeps (fuse bs r) = bs @ bmkeps r"
  using assms
  by (metis bmkeps_retrieve bnullable_correctness erase_fuse mkeps_nullable retrieve_fuse2)

lemma bmkepss_fuse: 
  assumes "bnullables rs"
  shows "bmkepss (map (fuse bs) rs) = bs @ bmkepss rs"
  using assms
  apply(induct rs arbitrary: bs)
  apply(auto simp add: bmkeps_fuse bnullable_fuse)
  done

lemma bder_fuse:
  shows "bder c (fuse bs a) = fuse bs  (bder c a)"
  apply(induct a arbitrary: bs c)
  apply(simp_all)
  done




inductive 
  rrewrite:: "arexp \<Rightarrow> arexp \<Rightarrow> bool" ("_ \<leadsto> _" [99, 99] 99)
and 
  srewrite:: "arexp list \<Rightarrow> arexp list \<Rightarrow> bool" (" _ s\<leadsto> _" [100, 100] 100)
where
  bs1: "ASEQ bs AZERO r2 \<leadsto> AZERO"
| bs2: "ASEQ bs r1 AZERO \<leadsto> AZERO"
| bs3: "ASEQ bs1 (AONE bs2) r \<leadsto> fuse (bs1@bs2) r"
| bs4: "r1 \<leadsto> r2 \<Longrightarrow> ASEQ bs r1 r3 \<leadsto> ASEQ bs r2 r3"
| bs5: "r3 \<leadsto> r4 \<Longrightarrow> ASEQ bs r1 r3 \<leadsto> ASEQ bs r1 r4"
| bs6: "AALTs bs [] \<leadsto> AZERO"
| bs7: "AALTs bs [r] \<leadsto> fuse bs r"
| bs8: "rs1 s\<leadsto> rs2 \<Longrightarrow> AALTs bs rs1 \<leadsto> AALTs bs rs2"
(*| ss1:  "[] s\<leadsto> []"*)
| ss2:  "rs1 s\<leadsto> rs2 \<Longrightarrow> (r # rs1) s\<leadsto> (r # rs2)"
| ss3:  "r1 \<leadsto> r2 \<Longrightarrow> (r1 # rs) s\<leadsto> (r2 # rs)"
| ss4:  "(AZERO # rs) s\<leadsto> rs"
| ss5:  "(AALTs bs1 rs1 # rsb) s\<leadsto> ((map (fuse bs1) rs1) @ rsb)"
| ss6:  "erase a1 = erase a2 \<Longrightarrow> (rsa@[a1]@rsb@[a2]@rsc) s\<leadsto> (rsa@[a1]@rsb@rsc)"


inductive 
  rrewrites:: "arexp \<Rightarrow> arexp \<Rightarrow> bool" ("_ \<leadsto>* _" [100, 100] 100)
where 
  rs1[intro, simp]:"r \<leadsto>* r"
| rs2[intro]: "\<lbrakk>r1 \<leadsto>* r2; r2 \<leadsto> r3\<rbrakk> \<Longrightarrow> r1 \<leadsto>* r3"

inductive 
  srewrites:: "arexp list \<Rightarrow> arexp list \<Rightarrow> bool" ("_ s\<leadsto>* _" [100, 100] 100)
where 
  sss1[intro, simp]:"rs s\<leadsto>* rs"
| sss2[intro]: "\<lbrakk>rs1 s\<leadsto> rs2; rs2 s\<leadsto>* rs3\<rbrakk> \<Longrightarrow> rs1 s\<leadsto>* rs3"


lemma r_in_rstar: 
  shows "r1 \<leadsto> r2 \<Longrightarrow> r1 \<leadsto>* r2"
  using rrewrites.intros(1) rrewrites.intros(2) by blast

lemma rrewrites_trans[trans]: 
  assumes a1: "r1 \<leadsto>* r2"  and a2: "r2 \<leadsto>* r3"
  shows "r1 \<leadsto>* r3"
  using a2 a1
  apply(induct r2 r3 arbitrary: r1 rule: rrewrites.induct) 
  apply(auto)
  done

lemma srewrites_trans[trans]: 
  assumes a1: "r1 s\<leadsto>* r2"  and a2: "r2 s\<leadsto>* r3"
  shows "r1 s\<leadsto>* r3"
  using a1 a2
  apply(induct r1 r2 arbitrary: r3 rule: srewrites.induct) 
   apply(auto)
  done



lemma contextrewrites0: 
  "rs1 s\<leadsto>* rs2 \<Longrightarrow> AALTs bs rs1 \<leadsto>* AALTs bs rs2"
  apply(induct rs1 rs2 rule: srewrites.inducts)
   apply simp
  using bs8 r_in_rstar rrewrites_trans by blast

lemma contextrewrites1: 
  "r \<leadsto>* r' \<Longrightarrow> AALTs bs (r # rs) \<leadsto>* AALTs bs (r' # rs)"
  apply(induct r r' rule: rrewrites.induct)
   apply simp
  using bs8 ss3 by blast

lemma srewrite1: 
  shows "rs1 s\<leadsto> rs2 \<Longrightarrow> (rs @ rs1) s\<leadsto> (rs @ rs2)"
  apply(induct rs)
   apply(auto)
  using ss2 by auto

lemma srewrites1: 
  shows "rs1 s\<leadsto>* rs2 \<Longrightarrow> (rs @ rs1) s\<leadsto>* (rs @ rs2)"
  apply(induct rs1 rs2 rule: srewrites.induct)
   apply(auto)
  using srewrite1 by blast

lemma srewrite2: 
  shows  "r1 \<leadsto> r2 \<Longrightarrow> True"
  and "rs1 s\<leadsto> rs2 \<Longrightarrow> (rs1 @ rs) s\<leadsto>* (rs2 @ rs)"
  apply(induct rule: rrewrite_srewrite.inducts)
  apply(auto)
  apply (metis append_Cons append_Nil srewrites1)
  apply(meson srewrites.simps ss3)
  apply (meson srewrites.simps ss4)
  apply (meson srewrites.simps ss5)
  by (metis append_Cons append_Nil srewrites.simps ss6)
  

lemma srewrites3: 
  shows "rs1 s\<leadsto>* rs2 \<Longrightarrow> (rs1 @ rs) s\<leadsto>* (rs2 @ rs)"
  apply(induct rs1 rs2 arbitrary: rs rule: srewrites.induct)
   apply(auto)
  by (meson srewrite2(2) srewrites_trans)

(*
lemma srewrites4:
  assumes "rs3 s\<leadsto>* rs4" "rs1 s\<leadsto>* rs2" 
  shows "(rs1 @ rs3) s\<leadsto>* (rs2 @ rs4)"
  using assms
  apply(induct rs3 rs4 arbitrary: rs1 rs2 rule: srewrites.induct)
  apply (simp add: srewrites3)
  using srewrite1 by blast
*)

lemma srewrites6:
  assumes "r1 \<leadsto>* r2" 
  shows "[r1] s\<leadsto>* [r2]"
  using assms
  apply(induct r1 r2 rule: rrewrites.induct)
  apply(auto)
  by (meson srewrites.simps srewrites_trans ss3)

lemma srewrites7:
  assumes "rs3 s\<leadsto>* rs4" "r1 \<leadsto>* r2" 
  shows "(r1 # rs3) s\<leadsto>* (r2 # rs4)"
  using assms
  by (smt (verit, del_insts) append.simps srewrites1 srewrites3 srewrites6 srewrites_trans)
  
lemma ss6_stronger_aux:
  shows "(rs1 @ rs2) s\<leadsto>* (rs1 @ distinctBy rs2 erase (set (map erase rs1)))"
  apply(induct rs2 arbitrary: rs1)
   apply(auto)
  apply (smt (verit, best) append.assoc append.right_neutral append_Cons append_Nil split_list srewrite2(2) srewrites_trans ss6)
  apply(drule_tac x="rs1 @ [a]" in meta_spec)
  apply(simp)
  done

lemma ss6_stronger:
  shows "rs1 s\<leadsto>* distinctBy rs1 erase {}"
  using ss6_stronger_aux[of "[]" _] by auto

lemma rewrite_preserves_fuse: 
  shows "r2 \<leadsto> r3 \<Longrightarrow> fuse bs r2 \<leadsto> fuse bs r3"
  and   "rs2 s\<leadsto> rs3 \<Longrightarrow> map (fuse bs) rs2 s\<leadsto> map (fuse bs) rs3"
proof(induct rule: rrewrite_srewrite.inducts)
  case (bs3 bs1 bs2 r)
  then show "fuse bs (ASEQ bs1 (AONE bs2) r) \<leadsto> fuse bs (fuse (bs1 @ bs2) r)"
    by (metis fuse.simps(5) fuse_append rrewrite_srewrite.bs3) 
next
  case (bs7 bs1 r)
  then show "fuse bs (AALTs bs1 [r]) \<leadsto> fuse bs (fuse bs1 r)"
    by (metis fuse.simps(4) fuse_append rrewrite_srewrite.bs7) 
next
  case (ss2 rs1 rs2 r)
  then show "map (fuse bs) (r # rs1) s\<leadsto> map (fuse bs) (r # rs2)"
    by (simp add: rrewrite_srewrite.ss2)
next
  case (ss3 r1 r2 rs)
  then show "map (fuse bs) (r1 # rs) s\<leadsto> map (fuse bs) (r2 # rs)"
    by (simp add: rrewrite_srewrite.ss3)
next
  case (ss5 bs1 rs1 rsb)
  have "map (fuse bs) (AALTs bs1 rs1 # rsb) = AALTs (bs @ bs1) rs1 # (map (fuse bs) rsb)" by simp
  also have "... s\<leadsto> ((map (fuse (bs @ bs1)) rs1) @ (map (fuse bs) rsb))"
    by (simp add: rrewrite_srewrite.ss5)  
  finally show "map (fuse bs) (AALTs bs1 rs1 # rsb) s\<leadsto> map (fuse bs) (map (fuse bs1) rs1 @ rsb)"
    by (simp add: comp_def fuse_append)
next
  case (ss6 a1 a2 rsa rsb rsc)
  then show "map (fuse bs) (rsa @ [a1] @ rsb @ [a2] @ rsc) s\<leadsto> map (fuse bs) (rsa @ [a1] @ rsb @ rsc)"
    apply(simp)
    apply(rule rrewrite_srewrite.ss6[simplified])
    apply(simp add: erase_fuse)
    done
qed (auto intro: rrewrite_srewrite.intros)

lemma rewrites_fuse:  
  assumes "r1 \<leadsto>* r2"
  shows "fuse bs r1 \<leadsto>* fuse bs r2"
using assms
apply(induction r1 r2 arbitrary: bs rule: rrewrites.induct)
apply(auto intro: rewrite_preserves_fuse)
done


lemma star_seq:  
  assumes "r1 \<leadsto>* r2"
  shows "ASEQ bs r1 r3 \<leadsto>* ASEQ bs r2 r3"
using assms
apply(induct r1 r2 arbitrary: r3 rule: rrewrites.induct)
apply(auto intro: rrewrite_srewrite.intros)
done

lemma star_seq2:  
  assumes "r3 \<leadsto>* r4"
  shows "ASEQ bs r1 r3 \<leadsto>* ASEQ bs r1 r4"
  using assms
apply(induct r3 r4 arbitrary: r1 rule: rrewrites.induct)
apply(auto intro: rrewrite_srewrite.intros)
done

lemma continuous_rewrite: 
  assumes "r1 \<leadsto>* AZERO"
  shows "ASEQ bs1 r1 r2 \<leadsto>* AZERO"
using assms bs1 star_seq by blast

(*
lemma continuous_rewrite2: 
  assumes "r1 \<leadsto>* AONE bs"
  shows "ASEQ bs1 r1 r2 \<leadsto>* (fuse (bs1 @ bs) r2)"
  using assms  by (meson bs3 rrewrites.simps star_seq)
*)

lemma bsimp_aalts_simpcases: 
  shows "AONE bs \<leadsto>* bsimp (AONE bs)"  
  and   "AZERO \<leadsto>* bsimp AZERO" 
  and   "ACHAR bs c \<leadsto>* bsimp (ACHAR bs c)"
  by (simp_all)

lemma bsimp_AALTs_rewrites: 
  shows "AALTs bs1 rs \<leadsto>* bsimp_AALTs bs1 rs"
  by (smt (verit) bs6 bs7 bsimp_AALTs.elims rrewrites.simps)

lemma trivialbsimp_srewrites: 
  assumes "\<And>x. x \<in> set rs \<Longrightarrow> x \<leadsto>* f x"
  shows "rs s\<leadsto>* (map f rs)"
using assms
  apply(induction rs)
  apply(simp_all add: srewrites7)
  done

lemma fltsfrewrites: "rs s\<leadsto>* flts rs"
  apply(induction rs rule: flts.induct)
  apply(auto intro: rrewrite_srewrite.intros)
  apply (meson srewrites.simps srewrites1 ss5)
  using rs1 srewrites7 apply presburger
  using srewrites7 apply force
  apply (simp add: srewrites7)
  by (simp add: srewrites7)

lemma bnullable0:
shows "r1 \<leadsto> r2 \<Longrightarrow> bnullable r1 = bnullable r2" 
  and "rs1 s\<leadsto> rs2 \<Longrightarrow> bnullables rs1 = bnullables rs2" 
  apply(induct rule: rrewrite_srewrite.inducts)
  apply(auto simp add:  bnullable_fuse)
  apply (meson UnCI bnullable_fuse imageI)
  by (metis bnullable_correctness)


lemma rewrites_bnullable_eq: 
  assumes "r1 \<leadsto>* r2" 
  shows "bnullable r1 = bnullable r2"
using assms 
  apply(induction r1 r2 rule: rrewrites.induct)
  apply simp
  using bnullable0(1) by auto

lemma rewrite_bmkeps_aux: 
  shows "r1 \<leadsto> r2 \<Longrightarrow> bnullable r1 \<Longrightarrow> bmkeps r1 = bmkeps r2"
  and   "rs1 s\<leadsto> rs2 \<Longrightarrow> bnullables rs1 \<Longrightarrow> bmkepss rs1 = bmkepss rs2" 
proof (induct rule: rrewrite_srewrite.inducts)
  case (bs3 bs1 bs2 r)
  have IH2: "bnullable (ASEQ bs1 (AONE bs2) r)" by fact
  then show "bmkeps (ASEQ bs1 (AONE bs2) r) = bmkeps (fuse (bs1 @ bs2) r)"
    by (simp add: bmkeps_fuse)
next
  case (bs7 bs r)
  have IH2: "bnullable (AALTs bs [r])" by fact
  then show "bmkeps (AALTs bs [r]) = bmkeps (fuse bs r)" 
    by (simp add: bmkeps_fuse)
next
  case (ss3 r1 r2 rs)
  have IH1: "bnullable r1 \<Longrightarrow> bmkeps r1 = bmkeps r2" by fact
  have as: "r1 \<leadsto> r2" by fact
  from IH1 as show "bmkepss (r1 # rs) = bmkepss (r2 # rs)"
    by (simp add: bnullable0)
next
  case (ss5 bs1 rs1 rsb)
  have "bnullables (AALTs bs1 rs1 # rsb)" by fact
  then show "bmkepss (AALTs bs1 rs1 # rsb) = bmkepss (map (fuse bs1) rs1 @ rsb)"
    by (simp add: bmkepss1 bmkepss2 bmkepss_fuse bnullable_fuse)
next
  case (ss6 a1 a2 rsa rsb rsc)
  have as1: "erase a1 = erase a2" by fact
  have as3: "bnullables (rsa @ [a1] @ rsb @ [a2] @ rsc)" by fact
  show "bmkepss (rsa @ [a1] @ rsb @ [a2] @ rsc) = bmkepss (rsa @ [a1] @ rsb @ rsc)" using as1 as3
    by (smt (verit, best) append_Cons bmkeps.simps(3) bmkepss.simps(2) bmkepss1 bmkepss2 bnullable_correctness)
qed (auto)

lemma rewrites_bmkeps: 
  assumes "r1 \<leadsto>* r2" "bnullable r1" 
  shows "bmkeps r1 = bmkeps r2"
  using assms
proof(induction r1 r2 rule: rrewrites.induct)
  case (rs1 r)
  then show "bmkeps r = bmkeps r" by simp
next
  case (rs2 r1 r2 r3)
  then have IH: "bmkeps r1 = bmkeps r2" by simp
  have a1: "bnullable r1" by fact
  have a2: "r1 \<leadsto>* r2" by fact
  have a3: "r2 \<leadsto> r3" by fact
  have a4: "bnullable r2" using a1 a2 by (simp add: rewrites_bnullable_eq) 
  then have "bmkeps r2 = bmkeps r3"
    using a3 bnullable0(1) rewrite_bmkeps_aux(1) by blast 
  then show "bmkeps r1 = bmkeps r3" using IH by simp
qed


lemma rewrites_to_bsimp: 
  shows "r \<leadsto>* bsimp r"
proof (induction r rule: bsimp.induct)
  case (1 bs1 r1 r2)
  have IH1: "r1 \<leadsto>* bsimp r1" by fact
  have IH2: "r2 \<leadsto>* bsimp r2" by fact
  { assume as: "bsimp r1 = AZERO \<or> bsimp r2 = AZERO"
    with IH1 IH2 have "r1 \<leadsto>* AZERO \<or> r2 \<leadsto>* AZERO" by auto
    then have "ASEQ bs1 r1 r2 \<leadsto>* AZERO"
      by (metis bs2 continuous_rewrite rrewrites.simps star_seq2)  
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 r1 r2)" using as by auto
  }
  moreover
  { assume "\<exists>bs. bsimp r1 = AONE bs"
    then obtain bs where as: "bsimp r1 = AONE bs" by blast
    with IH1 have "r1 \<leadsto>* AONE bs" by simp
    then have "ASEQ bs1 r1 r2 \<leadsto>* fuse (bs1 @ bs) r2" using bs3 star_seq by blast
    with IH2 have "ASEQ bs1 r1 r2 \<leadsto>* fuse (bs1 @ bs) (bsimp r2)"
      using rewrites_fuse by (meson rrewrites_trans) 
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 (AONE bs) r2)" by simp
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 r1 r2)" by (simp add: as) 
  } 
  moreover
  { assume as1: "bsimp r1 \<noteq> AZERO" "bsimp r2 \<noteq> AZERO" and as2: "(\<nexists>bs. bsimp r1 = AONE bs)" 
    then have "bsimp_ASEQ bs1 (bsimp r1) (bsimp r2) = ASEQ bs1 (bsimp r1) (bsimp r2)" 
      by (simp add: bsimp_ASEQ1) 
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp_ASEQ bs1 (bsimp r1) (bsimp r2)" using as1 as2 IH1 IH2
      by (metis rrewrites_trans star_seq star_seq2) 
    then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 r1 r2)" by simp
  } 
  ultimately show "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 r1 r2)" by blast
next
  case (2 bs1 rs)
  have IH: "\<And>x. x \<in> set rs \<Longrightarrow> x \<leadsto>* bsimp x" by fact
  then have "rs s\<leadsto>* (map bsimp rs)" by (simp add: trivialbsimp_srewrites)
  also have "... s\<leadsto>* flts (map bsimp rs)" by (simp add: fltsfrewrites) 
  also have "... s\<leadsto>* distinctBy (flts (map bsimp rs)) erase {}" by (simp add: ss6_stronger) 
  finally have "AALTs bs1 rs \<leadsto>* AALTs bs1 (distinctBy (flts (map bsimp rs)) erase {})"
    using contextrewrites0 by blast
  also have "... \<leadsto>* bsimp_AALTs  bs1 (distinctBy (flts (map bsimp rs)) erase {})"
    by (simp add: bsimp_AALTs_rewrites)     
  finally show "AALTs bs1 rs \<leadsto>* bsimp (AALTs bs1 rs)" by simp
qed (simp_all)


lemma to_zero_in_alt: 
  shows "AALT bs (ASEQ [] AZERO r) r2 \<leadsto> AALT bs AZERO r2"
  by (simp add: bs1 bs8 ss3)



lemma  bder_fuse_list: 
  shows "map (bder c \<circ> fuse bs1) rs1 = map (fuse bs1 \<circ> bder c) rs1"
  apply(induction rs1)
  apply(simp_all add: bder_fuse)
  done

lemma rewrite_preserves_bder: 
  shows "r1 \<leadsto> r2 \<Longrightarrow> bder c r1 \<leadsto>* bder c r2"
  and   "rs1 s\<leadsto> rs2 \<Longrightarrow> map (bder c) rs1 s\<leadsto>* map (bder c) rs2"
proof(induction rule: rrewrite_srewrite.inducts)
  case (bs1 bs r2)
  show "bder c (ASEQ bs AZERO r2) \<leadsto>* bder c AZERO"
    by (simp add: continuous_rewrite) 
next
  case (bs2 bs r1)
  show "bder c (ASEQ bs r1 AZERO) \<leadsto>* bder c AZERO"
    apply(auto)
    apply (meson bs6 contextrewrites0 rrewrite_srewrite.bs2 rs2 ss3 ss4 sss1 sss2)
    by (simp add: r_in_rstar rrewrite_srewrite.bs2)
next
  case (bs3 bs1 bs2 r)
  show "bder c (ASEQ bs1 (AONE bs2) r) \<leadsto>* bder c (fuse (bs1 @ bs2) r)"
    apply(simp)
    by (metis (no_types, lifting) bder_fuse bs8 bs7 fuse_append rrewrites.simps ss4 to_zero_in_alt)
next
  case (bs4 r1 r2 bs r3)
  have as: "r1 \<leadsto> r2" by fact
  have IH: "bder c r1 \<leadsto>* bder c r2" by fact
  from as IH show "bder c (ASEQ bs r1 r3) \<leadsto>* bder c (ASEQ bs r2 r3)"
    by (metis bder.simps(5) bnullable0(1) contextrewrites1 rewrite_bmkeps_aux(1) star_seq)
next
  case (bs5 r3 r4 bs r1)
  have as: "r3 \<leadsto> r4" by fact 
  have IH: "bder c r3 \<leadsto>* bder c r4" by fact 
  from as IH show "bder c (ASEQ bs r1 r3) \<leadsto>* bder c (ASEQ bs r1 r4)"
    apply(simp)
    apply(auto)
    using contextrewrites0 r_in_rstar rewrites_fuse srewrites6 srewrites7 star_seq2 apply presburger
    using star_seq2 by blast
next
  case (bs6 bs)
  show "bder c (AALTs bs []) \<leadsto>* bder c AZERO"
    using rrewrite_srewrite.bs6 by force 
next
  case (bs7 bs r)
  show "bder c (AALTs bs [r]) \<leadsto>* bder c (fuse bs r)"
    by (simp add: bder_fuse r_in_rstar rrewrite_srewrite.bs7) 
next
  case (bs8 rs1 rs2 bs)
  have IH1: "map (bder c) rs1 s\<leadsto>* map (bder c) rs2" by fact
  then show "bder c (AALTs bs rs1) \<leadsto>* bder c (AALTs bs rs2)" 
    using contextrewrites0 by force    
(*next
  case ss1
  show "map (bder c) [] s\<leadsto>* map (bder c) []" by simp*)
next
  case (ss2 rs1 rs2 r)
  have IH1: "map (bder c) rs1 s\<leadsto>* map (bder c) rs2" by fact
  then show "map (bder c) (r # rs1) s\<leadsto>* map (bder c) (r # rs2)"
    by (simp add: srewrites7) 
next
  case (ss3 r1 r2 rs)
  have IH: "bder c r1 \<leadsto>* bder c r2" by fact
  then show "map (bder c) (r1 # rs) s\<leadsto>* map (bder c) (r2 # rs)"
    by (simp add: srewrites7) 
next
  case (ss4 rs)
  show "map (bder c) (AZERO # rs) s\<leadsto>* map (bder c) rs"
    using rrewrite_srewrite.ss4 by fastforce 
next
  case (ss5 bs1 rs1 rsb)
  show "map (bder c) (AALTs bs1 rs1 # rsb) s\<leadsto>* map (bder c) (map (fuse bs1) rs1 @ rsb)"
    apply(simp)
    using bder_fuse_list map_map rrewrite_srewrite.ss5 srewrites.simps by blast
next
  case (ss6 a1 a2 bs rsa rsb)
  have as: "erase a1 = erase a2" by fact
  show "map (bder c) (bs @ [a1] @ rsa @ [a2] @ rsb) s\<leadsto>* map (bder c) (bs @ [a1] @ rsa @ rsb)"
    apply(simp only: map_append)
    by (smt (verit, best) erase_bder list.simps(8) list.simps(9) as rrewrite_srewrite.ss6 srewrites.simps)
qed

lemma rewrites_preserves_bder: 
  assumes "r1 \<leadsto>* r2"
  shows "bder c r1 \<leadsto>* bder c r2"
using assms  
apply(induction r1 r2 rule: rrewrites.induct)
apply(simp_all add: rewrite_preserves_bder rrewrites_trans)
done


lemma central:  
  shows "bders r s \<leadsto>* bders_simp r s"
proof(induct s arbitrary: r rule: rev_induct)
  case Nil
  then show "bders r [] \<leadsto>* bders_simp r []" by simp
next
  case (snoc x xs)
  have IH: "\<And>r. bders r xs \<leadsto>* bders_simp r xs" by fact
  have "bders r (xs @ [x]) = bders (bders r xs) [x]" by (simp add: bders_append)
  also have "... \<leadsto>* bders (bders_simp r xs) [x]" using IH
    by (simp add: rewrites_preserves_bder)
  also have "... \<leadsto>* bders_simp (bders_simp r xs) [x]" using IH
    by (simp add: rewrites_to_bsimp)
  finally show "bders r (xs @ [x]) \<leadsto>* bders_simp r (xs @ [x])" 
    by (simp add: bders_simp_append)
qed

lemma main_aux: 
  assumes "bnullable (bders r s)"
  shows "bmkeps (bders r s) = bmkeps (bders_simp r s)"
proof -
  have "bders r s \<leadsto>* bders_simp r s" by (rule central)
  then 
  show "bmkeps (bders r s) = bmkeps (bders_simp r s)" using assms
    by (rule rewrites_bmkeps)
qed  


theorem main_blexer_simp: 
  shows "blexer r s = blexer_simp r s"
  unfolding blexer_def blexer_simp_def
  by (metis central main_aux rewrites_bnullable_eq)


theorem blexersimp_correctness: 
  shows "lexer r s = blexer_simp r s"
  using blexer_correctness main_blexer_simp by simp


(* some tests *)

lemma asize_fuse:
  shows "asize (fuse bs r) = asize r"
  apply(induct r arbitrary: bs)
  apply(auto)
  done

lemma asize_rewrite2:
  shows "r1 \<leadsto> r2 \<Longrightarrow> asize r1 \<ge> asize r2"
  and   "rs1 s\<leadsto> rs2 \<Longrightarrow> (sum_list (map asize rs1)) \<ge> (sum_list (map asize rs2))"
   apply(induct rule: rrewrite_srewrite.inducts)
   apply(auto simp add: asize_fuse comp_def)
  done

lemma asize_rrewrites:
  assumes "r1 \<leadsto>* r2"
  shows "asize r1 \<ge> asize r2"
  using assms
  apply(induct rule: rrewrites.induct)
   apply(auto)
  using asize_rewrite2(1) le_trans by blast
  


fun asize2 :: "arexp \<Rightarrow> nat" where
  "asize2 AZERO = 1"
| "asize2 (AONE cs) = 1" 
| "asize2 (ACHAR cs c) = 1"
| "asize2 (AALTs cs rs) = Suc (Suc (sum_list (map asize2 rs)))"
| "asize2 (ASEQ cs r1 r2) = Suc (asize2 r1 + asize2 r2)"
| "asize2 (ASTAR cs r) = Suc (asize2 r)"


lemma asize2_fuse:
  shows "asize2 (fuse bs r) = asize2 r"
  apply(induct r arbitrary: bs)
  apply(auto)
  done

lemma asize2_not_zero:
  shows "0 < asize2 r"
  apply(induct r)
  apply(auto)
  done

lemma asize_rewrite:
  shows "r1 \<leadsto> r2 \<Longrightarrow> asize2 r1 > asize2 r2"
  and   "rs1 s\<leadsto> rs2 \<Longrightarrow> (sum_list (map asize2 rs1)) > (sum_list (map asize2 rs2))"
   apply(induct rule: rrewrite_srewrite.inducts)
   apply(auto simp add: asize2_fuse comp_def)
  apply(simp add: asize2_not_zero) 
  done

lemma asize2_bsimp_ASEQ:
  shows "asize2 (bsimp_ASEQ bs r1 r2) \<le> Suc (asize2 r1 + asize2 r2)"
  apply(induct bs r1 r2 rule: bsimp_ASEQ.induct)
  apply(auto)
  done

lemma asize2_bsimp_AALTs:
  shows "asize2 (bsimp_AALTs bs rs) \<le> Suc (Suc (sum_list (map asize2 rs)))"
  apply(induct bs rs rule: bsimp_AALTs.induct)
  apply(auto simp add: asize2_fuse)
  done

lemma distinctBy_asize2:
  shows "sum_list (map asize2 (distinctBy rs f acc)) \<le> sum_list (map asize2 rs)"
  apply(induct rs f acc rule: distinctBy.induct)
  apply(auto)
  done

lemma flts_asize2:
  shows "sum_list (map asize2 (flts rs)) \<le> sum_list (map asize2 rs)"
  apply(induct rs rule: flts.induct)
  apply(auto simp add: comp_def asize2_fuse)
  done

lemma sumlist_asize2:
  assumes "\<And>x. x \<in> set rs \<Longrightarrow> asize2 (f x) \<le> asize2 x"
  shows "sum_list (map asize2 (map f rs)) \<le> sum_list (map asize2 rs)"
  using assms
  apply(induct rs)
   apply(auto simp add: comp_def)
  by (simp add: add_le_mono)

lemma test0:
  assumes "r1 \<leadsto>* r2"
  shows "r1 = r2 \<or> (\<exists>r3. r1 \<leadsto> r3 \<and> r3 \<leadsto>* r2)"
  using assms
  apply(induct r1 r2 rule: rrewrites.induct)
   apply(auto)
  done

lemma test2:
  assumes "r1 \<leadsto>* r2"
  shows "asize2 r1 \<ge> asize2 r2"
using assms
  apply(induct r1 r2 rule: rrewrites.induct)
  apply(auto)
  using asize_rewrite(1) by fastforce
  

lemma test3:
  shows "r = bsimp r \<or> (asize2 (bsimp r) < asize2 r)"
proof -
  have "r \<leadsto>* bsimp r"
    by (simp add: rewrites_to_bsimp)
  then have "r = bsimp r \<or> (\<exists>r3. r \<leadsto> r3 \<and> r3 \<leadsto>* bsimp r)"
    using test0 by blast
  then show ?thesis
    by (meson asize_rewrite(1) dual_order.strict_trans2 test2)
qed

lemma test3Q:
  shows "r = bsimp r \<or> (asize (bsimp r) \<le> asize r)"
proof -
  have "r \<leadsto>* bsimp r"
    by (simp add: rewrites_to_bsimp)
  then have "r = bsimp r \<or> (\<exists>r3. r \<leadsto> r3 \<and> r3 \<leadsto>* bsimp r)"
    using test0 by blast
  then show ?thesis
    using asize_rewrite2(1) asize_rrewrites le_trans by blast
qed

lemma test4:
  shows "asize2 (bsimp (bsimp r)) \<le> asize2 (bsimp r)"
  apply(induct r rule: bsimp.induct)
       apply(auto)
  using rewrites_to_bsimp test2 apply fastforce
  using rewrites_to_bsimp test2 by presburger

lemma test4Q:
  shows "asize (bsimp (bsimp r)) \<le> asize (bsimp r)"
  apply(induct r rule: bsimp.induct)
       apply(auto)
  apply (metis order_refl test3Q)
  by (metis le_refl test3Q)



lemma testb0:
  shows "fuse bs1 (bsimp_ASEQ bs r1 r2) =  bsimp_ASEQ (bs1 @ bs) r1 r2"
  apply(induct bs r1 r2 arbitrary: bs1 rule: bsimp_ASEQ.induct)
  apply(auto)
  done

lemma testb1:
  shows "fuse bs1 (bsimp_AALTs bs rs) =  bsimp_AALTs (bs1 @ bs) rs"
  apply(induct bs rs arbitrary: bs1 rule: bsimp_AALTs.induct)
  apply(auto simp add: fuse_append)
  done

lemma testb2:
  shows "bsimp (bsimp_ASEQ bs r1 r2) =  bsimp_ASEQ bs (bsimp r1) (bsimp r2)"
  apply(induct bs r1 r2 rule: bsimp_ASEQ.induct)
  apply(auto simp add: testb0 testb1)
  done  

lemma testb3:
  shows "\<nexists>r'. (bsimp r \<leadsto> r') \<and> asize2 (bsimp r) > asize2 r'"
apply(induct r rule: bsimp.induct)
       apply(auto)
       defer
       defer
  using rrewrite.cases apply blast
  using rrewrite.cases apply blast
  using rrewrite.cases apply blast
  using rrewrite.cases apply blast
  oops

lemma testb4:
  assumes "sum_list (map asize rs1) \<le> sum_list (map asize rs2)"
  shows "asize (bsimp_AALTs bs1 rs1) \<le> Suc (asize (bsimp_AALTs bs1 rs2))"
  using assms 
apply(induct bs1 rs2 arbitrary: rs1 rule: bsimp_AALTs.induct)
    apply(auto)
    apply(case_tac rs1)
     apply(auto)
  using asize2.elims apply auto[1]
  apply (metis One_nat_def Zero_not_Suc asize.elims)
      apply(case_tac rs1)
    apply(auto)
   apply(case_tac list)
    apply(auto)
  using asize_fuse apply force
  apply (simp add: asize_fuse)
  by (smt (verit, ccfv_threshold) One_nat_def add.right_neutral asize.simps(1) asize.simps(4) asize_fuse bsimp_AALTs.elims le_Suc_eq list.map(1) list.map(2) not_less_eq_eq sum_list_simps(1) sum_list_simps(2))

lemma flts_asize:
  shows "sum_list (map asize (flts rs)) \<le> sum_list (map asize rs)"
  apply(induct rs rule: flts.induct)
  apply(auto simp add: comp_def asize_fuse)
  done


lemma test5:
  shows "asize2 r \<ge> asize2 (bsimp r)"
  apply(induct r rule: bsimp.induct)
       apply(auto)
  apply (meson Suc_le_mono add_le_mono asize2_bsimp_ASEQ order_trans)
  apply(rule order_trans)
   apply(rule asize2_bsimp_AALTs)
  apply(simp)
  apply(rule order_trans)
   apply(rule distinctBy_asize2)
  apply(rule order_trans)
   apply(rule flts_asize2)
  using sumlist_asize2 by force
  

fun awidth :: "arexp \<Rightarrow> nat" where
  "awidth AZERO = 1"
| "awidth (AONE cs) = 1" 
| "awidth (ACHAR cs c) = 1"
| "awidth (AALTs cs rs) = (sum_list (map awidth rs))"
| "awidth (ASEQ cs r1 r2) = (awidth r1 + awidth r2)"
| "awidth (ASTAR cs r) = (awidth r)"



lemma 
  shows "s \<notin> L r \<Longrightarrow> blexer_simp r s = None"
  by (simp add: blexersimp_correctness lexer_correct_None)

lemma g1:
  "bders_simp AZERO s = AZERO"
 apply(induct s)
 apply(simp)
 apply(simp)
  done

lemma g2:
  "s \<noteq> Nil \<Longrightarrow> bders_simp (AONE bs) s = AZERO"
 apply(induct s)
 apply(simp)
  apply(simp)
  apply(case_tac s)
  apply(simp)
       apply(simp)
  done

lemma finite_pder:
  shows "finite (pder c r)"
  apply(induct c r rule: pder.induct)
  apply(auto)
  done



lemma awidth_fuse:
  shows "awidth (fuse bs r) = awidth r"
  apply(induct r arbitrary: bs)
  apply(auto)
  done

lemma pders_SEQs:
  assumes "finite A"
  shows "card (SEQs A (STAR r)) \<le> card A"
  using assms
  by (simp add: SEQs_eq_image card_image_le)

lemma binullable_intern:
  shows "bnullable (intern r) = nullable r"
  apply(induct r)
  apply(auto simp add: bnullable_fuse)
  done

lemma
  "card (pder c r) \<le> awidth (bder c (intern r))"
  apply(induct c r rule: pder.induct)
  apply(simp)
      apply(simp)
     apply(simp)
    apply(simp)
  apply(rule order_trans)
     apply(rule card_Un_le)
  apply (simp add: awidth_fuse bder_fuse)
   defer
   apply(simp)
  apply(rule order_trans)
    apply(rule pders_SEQs)
  using finite_pder apply presburger
  apply (simp add: awidth_fuse)
  apply(auto)
     apply(rule order_trans)
      apply(rule card_Un_le)
     apply(simp add: awidth_fuse)
     defer
  using binullable_intern apply blast
  using binullable_intern apply blast
  apply (smt (verit, best) SEQs_eq_image add.commute add_Suc_right card_image_le dual_order.trans finite_pder trans_le_add2)
  apply(subgoal_tac "card (SEQs (pder c r1) r2) \<le> card (pder c r1)")
  apply(linarith)
    by (simp add: UNION_singleton_eq_range card_image_le finite_pder)
  
lemma
  "card (pder c r) \<le> asize (bder c (intern r))"
  apply(induct c r rule: pder.induct)
       apply(simp)
  apply(simp)
     apply(simp)
    apply(simp)
  apply (metis add_mono_thms_linordered_semiring(1) asize_fuse bder_fuse card_Un_le le_Suc_eq order_trans)
  defer 
   apply(simp)
   apply(rule order_trans)
    apply(rule pders_SEQs)
  using finite_pder apply presburger
  apply (simp add: asize_fuse)
  apply(simp)
  apply(auto)
  apply(rule order_trans)
      apply(rule card_Un_le)
  apply (smt (z3) SEQs_eq_image add.commute add_Suc_right add_mono_thms_linordered_semiring(1) asize_fuse card_image_le dual_order.trans finite_pder le_add1)
    apply(rule order_trans)
     apply(rule card_Un_le)
  using binullable_intern apply blast
  using binullable_intern apply blast
  by (smt (verit, best) SEQs_eq_image add.commute add_Suc_right card_image_le dual_order.trans finite_pder trans_le_add2)
  
lemma
  "card (pder c r) \<le> asize (bsimp (bder c (intern r)))"
  apply(induct c r rule: pder.induct)
       apply(simp)
      apply(simp)
     apply(simp)
    apply(simp)
  apply(rule order_trans)
     apply(rule card_Un_le)
    prefer 3
    apply(simp)
  apply(rule order_trans)
     apply(rule pders_SEQs)
  using finite_pder apply blast
  oops
  

(* below is the idempotency of bsimp *)

lemma bsimp_ASEQ_fuse:
  shows "fuse bs1 (bsimp_ASEQ bs2 r1 r2) = bsimp_ASEQ (bs1 @ bs2) r1 r2"
  apply(induct r1 r2 arbitrary: bs1 bs2 rule: bsimp_ASEQ.induct)
  apply(auto)
  done

lemma bsimp_AALTs_fuse:
  assumes "\<forall>r \<in> set rs. fuse bs1 (fuse bs2 r) = fuse (bs1 @ bs2) r"
  shows "fuse bs1 (bsimp_AALTs bs2 rs) = bsimp_AALTs (bs1 @ bs2) rs"
  using assms
  apply(induct bs2 rs arbitrary: bs1 rule: bsimp_AALTs.induct)
  apply(auto)
  done

lemma bsimp_fuse:
  shows "fuse bs (bsimp r) = bsimp (fuse bs r)"
  apply(induct r arbitrary: bs)
  apply(simp_all add: bsimp_ASEQ_fuse bsimp_AALTs_fuse fuse_append)
  done

lemma bsimp_ASEQ_idem:
  assumes "bsimp (bsimp r1) = bsimp r1" "bsimp (bsimp r2) = bsimp r2"
  shows "bsimp (bsimp_ASEQ x1 (bsimp r1) (bsimp r2)) = bsimp_ASEQ x1 (bsimp r1) (bsimp r2)"
  using assms
  apply(case_tac "bsimp r1 = AZERO")
  apply(simp)
  apply(case_tac "bsimp r2 = AZERO")
  apply(simp)
  apply(case_tac "\<exists>bs. bsimp r1 = AONE bs")
  apply(auto)[1]
  apply (metis bsimp_fuse)
  apply(simp add: bsimp_ASEQ1)
  done  

lemma bsimp_AALTs_idem:
  assumes "\<forall>r \<in> set rs. bsimp (bsimp r) = bsimp r" 
  shows "bsimp (bsimp_AALTs bs rs) = bsimp_AALTs bs (map bsimp rs)" 
  using assms
  apply(induct bs rs rule: bsimp_AALTs.induct)
  apply(simp)
   apply(simp)
  using bsimp_fuse apply presburger
  oops   
  
lemma bsimp_idem_rev:
  shows "\<nexists>r2. bsimp r1 \<leadsto> r2"
  apply(induct r1 rule: bsimp.induct)
  apply(auto)
  defer
  defer
  using rrewrite.simps apply blast
  using rrewrite.cases apply blast
  using rrewrite.simps apply blast
  using rrewrite.cases apply blast
  apply(case_tac "bsimp r1 = AZERO")
  apply(simp)
  apply(case_tac "bsimp r2 = AZERO")
  apply(simp)
  apply(case_tac "\<exists>bs. bsimp r1 = AONE bs")
  apply(auto)[1]
  prefer 2
  apply (smt (verit, best) arexp.distinct(25) arexp.inject(3) bsimp_ASEQ1 rrewrite.simps)
  defer
  oops

lemma bsimp_idem:
  shows "bsimp (bsimp r) = bsimp r"
  apply(induct r rule: bsimp.induct)
  apply(auto)
  using bsimp_ASEQ_idem apply presburger
  sorry


lemma neg:
  shows " \<not>(\<exists>r2. r1 \<leadsto> r2 \<and>  (r2 \<leadsto>* bsimp r1) )"
  apply(rule notI)
  apply(erule exE)
  apply(erule conjE)
  oops



lemma reduction_always_in_bsimp:
  shows " \<lbrakk> r1 \<leadsto> r2 ; \<not>(r2 \<leadsto>* bsimp r1)\<rbrakk> \<Longrightarrow> False"
  apply(erule rrewrite.cases)
         apply simp
        apply auto

  oops



(*
AALTs [] [AZERO, AALTs(bs1, [a, b]) ] 
rewrite seq 1: \<leadsto> AALTs [] [ AALTs(bs1, [a, b]) ] \<leadsto>
fuse [] (AALTs bs1, [a, b])
rewrite seq 2: \<leadsto> AALTs [] [AZERO, (fuse bs1 a), (fuse bs1 b)]) ]

*)

lemma normal_bsimp: 
  shows "\<nexists>r'. bsimp r \<leadsto> r'"
  oops

  (*r' size bsimp r > size r' 
    r' \<leadsto>* bsimp bsimp r
size bsimp r > size r' \<ge> size bsimp bsimp r*)

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 
  rerase :: "arexp \<Rightarrow> rrexp"
where
  "rerase AZERO = RZERO"
| "rerase (AONE _) = RONE"
| "rerase (ACHAR _ c) = RCHAR c"
| "rerase (AALTs bs rs) = RALTS (map rerase rs)"
| "rerase (ASEQ _ r1 r2) = RSEQ (rerase r1) (rerase r2)"
| "rerase (ASTAR _ r) = RSTAR (rerase r)"




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)"




lemma rsimp_aalts_smaller:
  shows "rsize (rsimp_ALTs  rs) \<le> rsize (RALTS rs)"
  apply(induct rs)
   apply simp
  sorry

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



lemma rerase_bsimp:
  shows "rerase (bsimp r) = rsimp (rerase r)"
  apply(induct r)
       apply auto


  sorry


lemma rerase_bder:
  shows "rerase (bder c r) = rder c (rerase r)"
  apply(induct r)
       apply auto
  sorry
(*
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 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 [] = []"

thm rsimp_SEQ.simps

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

lemma no_dup_after_simp:
  shows "RALTS rs = rsimp r \<Longrightarrow> distinct rs"
  sorry

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

lemma longlist_withstands_rsimp_alts:
  shows "length rs \<ge> 2 \<Longrightarrow> rsimp_ALTs rs = RALTS rs"
  sorry

lemma no_alt_short_list_after_simp:
  shows "RALTS rs = rsimp r \<Longrightarrow> rsimp_ALTs rs = RALTS 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 " rerase  (bders_simp r [c]) = rerase (bsimp (bders 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
  apply(metis idiot2)
  done

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




(*this section deals with the property of distinctBy: creates a list without duplicates*)

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 r s) = (rsimp (rders r s))"
and "\<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))) )"
  apply(induct s arbitrary: r r1 r2 rule: rev_induct)
     apply simp
  apply simp
   apply(case_tac "xs = []")

   apply (simp add: bders_simp_append )
   apply(subgoal_tac "rerase (bsimp (bder x (bders_simp r xs))) = (rsimp (rerase (bder x (bders_simp r xs)))) ")
    prefer 2
    apply (simp add: rerase_bsimp)
   apply(subgoal_tac "(rsimp (rerase (bder x (bders_simp r xs)))) = (rsimp (rder x (rerase (bders_simp r xs))))")
   
    apply(subgoal_tac "xs \<noteq> [] \<Longrightarrow> rsimp (rder x (rerase (bders_simp r xs))) = rsimp (rder x (rerase (bsimp (bders r xs))))")
  prefer 2
     apply presburger
  apply(case_tac "xs = []")
  sorry


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]"


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 finite_width_alt:
  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 r1 s) r2) #
           (rders_cond_list r2 (nullable_bools r1 (orderedPref s))  (orderedSuf s))) )) < N3"

  sorry


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

lemma finite_seq:
  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 finite_width_alt)
  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 (simp add: less_SucI shape_derssimp_seq(2))
   apply (meson less_SucI less_max_iff_disj)
  apply simp
  done


(*For star related error bound*)

lemma star_is_a_singleton_list_derc:
  shows " \<exists>Ss.  rders_simp (RSTAR r) [c] = rsimp_ALTs (map (\<lambda>s1.  rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)"
  apply simp
  apply(rule_tac x = "[[c]]" in exI)
  apply auto
  done

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

lemma double_nested_ALTs_under_rsimp:
  shows "rsimp (rsimp_ALTs ((RALTS rs1) # rs)) = rsimp (RALTS (rs1 @ rs))"
  apply(case_tac rs1)
  apply simp
  
   apply (metis list.simps(8) list.simps(9) neq_Nil_conv rdistinct.simps(1) rflts.simps(1) rflts.simps(2) rsimp.simps(2) rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3))
  apply(case_tac rs)
   apply simp
  apply auto
  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 der_seqstar_res:
  shows "rder x (RSEQ r1 r2) = RSEQ r3 r4"


lemma linearity_of_list_of_star_or_starseqs:
  shows "\<exists>Ssa. rsimp (rder x (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))) =
                 rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ssa)"
  apply(simp add: rder_rsimp_ALTs_commute)
  apply(induct Ss)
   apply simp
   apply (metis list.simps(8) rsimp_ALTs.simps(1))


  sorry

lemma starder_is_a_list_of_stars_or_starseqs:
  shows "s \<noteq> [] \<Longrightarrow> \<exists>Ss.  rders_simp (RSTAR r) s = rsimp_ALTs (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 auto
  apply(simp add: rders_simp_append)
  using linearity_of_list_of_star_or_starseqs 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"

  sorry


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:
  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 finite_seq apply blast
   prefer 2

   apply (simp add: finite_star)
  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"
*)


inductive aggressive:: "arexp \<Rightarrow> arexp \<Rightarrow> bool" ("_ \<leadsto>? _" [99, 99] 99)
  where
 "ASEQ bs (AALTs bs1 rs) r \<leadsto>? AALTs (bs@bs1) (map (\<lambda>r'. ASEQ [] r' r) rs) "



end