diff -r 549257d0b8b2 -r 8194086c2a8a thys2/SizeBound2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/SizeBound2.thy Thu Jan 20 01:48:18 2022 +0000 @@ -0,0 +1,1434 @@ + +theory SizeBound2 + imports "Lexer" +begin + +section \Bit-Encodings\ + +datatype bit = Z | S + +fun code :: "val \ 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 \ val \ val" +where + "Stars_add v (Stars vs) = Stars (v # vs)" + +function + decode' :: "bit list \ rexp \ (val * bit list)" +where + "decode' ds ZERO = (Void, [])" +| "decode' ds ONE = (Void, ds)" +| "decode' ds (CH d) = (Char d, ds)" +| "decode' [] (ALT r1 r2) = (Void, [])" +| "decode' (Z # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r1 in (Left v, ds'))" +| "decode' (S # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r2 in (Right v, ds'))" +| "decode' ds (SEQ r1 r2) = (let (v1, ds') = decode' ds r1 in + let (v2, ds'') = decode' ds' r2 in (Seq v1 v2, ds''))" +| "decode' [] (STAR r) = (Void, [])" +| "decode' (S # ds) (STAR r) = (Stars [], ds)" +| "decode' (Z # ds) (STAR r) = (let (v, ds') = decode' ds r in + let (vs, ds'') = decode' ds' (STAR r) + in (Stars_add v vs, ds''))" +by pat_completeness auto + +lemma decode'_smaller: + assumes "decode'_dom (ds, r)" + shows "length (snd (decode' ds r)) \ length ds" +using assms +apply(induct ds 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 \ rexp \ val option" +where + "decode ds r \ (let (v, ds') = decode' ds r + in (if ds' = [] then Some v else None))" + +lemma decode'_code_Stars: + assumes "\v\set vs. \ v : r \ (\x. decode' (code v @ x) r = (v, x)) \ flat v \ []" + shows "decode' (code (Stars vs) @ ds) (STAR r) = (Stars vs, ds)" + using assms + apply(induct vs) + apply(auto) + done + +lemma decode'_code: + assumes "\ 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 "\ 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 \ AALTs bs [r1, r2]" + +fun asize :: "arexp \ 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 \ 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 \ arexp \ 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 + +lemma fuse_Nil: + shows "fuse [] r = r" + by (induct r)(simp_all) + +lemma map_fuse_Nil: + shows "map (fuse []) rs = rs" + by (induct rs)(simp_all add: fuse_Nil) + + +fun intern :: "rexp \ 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 \ val \ 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 \ bool" +where + "bnullable (AZERO) = False" +| "bnullable (AONE bs) = True" +| "bnullable (ACHAR bs c) = False" +| "bnullable (AALTs bs rs) = (\r \ set rs. bnullable r)" +| "bnullable (ASEQ bs r1 r2) = (bnullable r1 \ bnullable r2)" +| "bnullable (ASTAR bs r) = True" + +fun + bmkeps :: "arexp \ bit list" +where + "bmkeps(AONE bs) = bs" +| "bmkeps(ASEQ bs r1 r2) = bs @ (bmkeps r1) @ (bmkeps r2)" +| "bmkeps(AALTs bs [r]) = bs @ (bmkeps r)" +| "bmkeps(AALTs bs (r#rs)) = (if bnullable(r) then bs @ (bmkeps r) else (bmkeps (AALTs bs rs)))" +| "bmkeps(ASTAR bs r) = bs @ [S]" + + +fun + bder :: "char \ arexp \ 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 \ string \ arexp" +where + "bders r [] = r" +| "bders r (c#s) = bders (bder c r) s" + +lemma bders_append: + "bders r (s1 @ s2) = bders (bders r s1) s2" + apply(induct s1 arbitrary: r 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 "\v\set vs. \ v : r \ 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 "\ 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] + defer + using retrieve_encode_STARS + apply(auto elim!: Prf_elims)[1] + apply(case_tac vs) + apply(simp) + apply(simp) + (* AALTs case *) + apply(simp) + apply(case_tac x2a) + apply(simp) + apply(auto elim!: Prf_elims)[1] + apply(simp) + apply(case_tac list) + apply(simp) + apply(auto) + apply(auto elim!: Prf_elims)[1] + done + +lemma retrieve_fuse: + assumes "\ 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 "\ 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 bnullable_Hdbmkeps_Hd: + assumes "bnullable a" + shows "bmkeps (AALTs bs (a # rs)) = bs @ (bmkeps a)" + using assms + by (metis bmkeps.simps(3) bmkeps.simps(4) list.exhaust) + +lemma r1: + assumes "\ bnullable a" "bnullable (AALTs bs rs)" + shows "bmkeps (AALTs bs (a # rs)) = bmkeps (AALTs bs rs)" + using assms + apply(induct rs) + apply(auto) + done + +lemma r2: + assumes "x \ set rs" "bnullable x" + shows "bnullable (AALTs bs rs)" + using assms + apply(induct rs) + apply(auto) + done + +lemma r3: + assumes "\ bnullable r" + " \ x \ set rs. bnullable x" + shows "retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs))) = + retrieve (AALTs bs (r # rs)) (mkeps (erase (AALTs bs (r # rs))))" + using assms + apply(induct rs arbitrary: r bs) + apply(auto)[1] + apply(auto) + using bnullable_correctness apply blast + apply(auto simp add: bnullable_correctness mkeps_nullable retrieve_fuse2) + apply(subst retrieve_fuse2[symmetric]) + apply (smt bnullable.simps(4) bnullable_correctness erase.simps(5) erase.simps(6) insert_iff list.exhaust list.set(2) mkeps.simps(3) mkeps_nullable) + apply(simp) + apply(case_tac "bnullable a") + apply (smt append_Nil2 bnullable.simps(4) bnullable_correctness erase.simps(5) erase.simps(6) fuse.simps(4) insert_iff list.exhaust list.set(2) mkeps.simps(3) mkeps_nullable retrieve_fuse2) + apply(drule_tac x="a" in meta_spec) + apply(drule_tac x="bs" in meta_spec) + apply(drule meta_mp) + apply(simp) + apply(drule meta_mp) + apply(auto) + apply(subst retrieve_fuse2[symmetric]) + apply(case_tac rs) + apply(simp) + apply(auto)[1] + apply (simp add: bnullable_correctness) + apply (metis append_Nil2 bnullable_correctness erase_fuse fuse.simps(4) list.set_intros(1) mkeps.simps(3) mkeps_nullable nullable.simps(4) r2) + apply (simp add: bnullable_correctness) + apply (metis append_Nil2 bnullable_correctness erase.simps(6) erase_fuse fuse.simps(4) list.set_intros(2) mkeps.simps(3) mkeps_nullable r2) + apply(simp) + done + + +lemma t: + assumes "\r \ set rs. nullable (erase r) \ bmkeps r = retrieve r (mkeps (erase r))" + "nullable (erase (AALTs bs rs))" + shows " bmkeps (AALTs bs rs) = retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs)))" + using assms + apply(induct rs arbitrary: bs) + apply(simp) + apply(auto simp add: bnullable_correctness) + apply(case_tac rs) + apply(auto simp add: bnullable_correctness)[2] + apply(subst r1) + apply(simp) + apply(rule r2) + apply(assumption) + apply(simp) + apply(drule_tac x="bs" in meta_spec) + apply(drule meta_mp) + apply(auto)[1] + prefer 2 + apply(case_tac "bnullable a") + apply(subst bnullable_Hdbmkeps_Hd) + apply blast + apply(subgoal_tac "nullable (erase a)") + prefer 2 + using bnullable_correctness apply blast + apply (metis (no_types, lifting) erase.simps(5) erase.simps(6) list.exhaust mkeps.simps(3) retrieve.simps(3) retrieve.simps(4)) + apply(subst r1) + apply(simp) + using r2 apply blast + apply(drule_tac x="bs" in meta_spec) + apply(drule meta_mp) + apply(auto)[1] + apply(simp) + using r3 apply blast + apply(auto) + using r3 by blast + +lemma bmkeps_retrieve: + assumes "nullable (erase r)" + shows "bmkeps r = retrieve r (mkeps (erase r))" + using assms + apply(induct r) + apply(simp) + apply(simp) + apply(simp) + apply(simp) + defer + apply(simp) + apply(rule t) + apply(auto) + done + +lemma bder_retrieve: + assumes "\ 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) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(case_tac "c = ca") + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(simp) + apply(erule Prf_elims) + 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) + apply (smt Prf_elims(3) injval.simps(2) injval.simps(3) retrieve.simps(4) retrieve.simps(5) same_append_eq) + apply(simp) + apply(case_tac "nullable (erase r1)") + apply(simp) + apply(erule Prf_elims) + apply(subgoal_tac "bnullable r1") + prefer 2 + using bnullable_correctness apply blast + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(subgoal_tac "bnullable r1") + prefer 2 + using bnullable_correctness apply blast + apply(simp) + apply(simp add: retrieve_fuse2) + apply(simp add: bmkeps_retrieve) + apply(simp) + apply(erule Prf_elims) + apply(simp) + using bnullable_correctness apply blast + apply(rename_tac bs r v) + apply(simp) + apply(erule Prf_elims) + apply(clarify) + apply(erule Prf_elims) + apply(clarify) + apply(subst injval.simps) + apply(simp del: retrieve.simps) + apply(subst retrieve.simps) + apply(subst retrieve.simps) + apply(simp) + apply(simp add: retrieve_fuse2) + done + + + +lemma MAIN_decode: + assumes "\ 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 "\ v : ders [] r" by fact + then have "\ 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: "\v. \ v : ders s r \ + Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r" by fact + have asm: "\ v : ders (s @ [c]) r" by fact + then have asm2: "\ 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 blex where + "blex a s \ if bnullable (bders a s) then Some (bmkeps (bders a s)) else None" + + + +definition blexer where + "blexer r s \ 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 \ bders (intern r) s" + define ds where "ds \ ders s r" + assume asm: "nullable ds" + have era: "erase bds = ds" + unfolding ds_def bds_def by simp + have mke: "\ 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 by (simp add: bmkeps_retrieve) + 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 + apply(subst bnullable_correctness[symmetric]) + apply(simp) + done +qed + + +fun distinctBy :: "'a list \ ('a \ 'b) \ 'b set \ 'a list" + where + "distinctBy [] f acc = []" +| "distinctBy (x#xs) f acc = + (if (f x) \ acc then distinctBy xs f acc + else x # (distinctBy xs f ({f x} \ acc)))" + +lemma dB_single_step: + shows "distinctBy (a#rs) f {} = a # distinctBy rs f {f a}" + by simp + +fun flts :: "arexp list \ 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 \ arexp \ arexp \ 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" + + +fun bsimp_AALTs :: "bit list \ arexp list \ arexp" + where + "bsimp_AALTs _ [] = AZERO" +| "bsimp_AALTs bs1 [r] = fuse bs1 r" +| "bsimp_AALTs bs1 rs = AALTs bs1 rs" + + +fun bsimp :: "arexp \ 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 \ string \ 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 \ if bnullable (bders_simp (intern r) s) then + decode (bmkeps (bders_simp (intern r) s)) r else None" + +export_code bders_simp in Scala module_name Example + +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 L_bsimp_ASEQ: + "L (SEQ (erase r1) (erase r2)) = L (erase (bsimp_ASEQ bs r1 r2))" + apply(induct bs r1 r2 rule: bsimp_ASEQ.induct) + apply(simp_all) + by (metis erase_fuse fuse.simps(4)) + +lemma L_bsimp_AALTs: + "L (erase (AALTs bs rs)) = L (erase (bsimp_AALTs bs rs))" + apply(induct bs rs rule: bsimp_AALTs.induct) + apply(simp_all add: erase_fuse) + done + +lemma L_erase_AALTs: + shows "L (erase (AALTs bs rs)) = \ (L ` erase ` (set rs))" + apply(induct rs) + apply(simp) + apply(simp) + apply(case_tac rs) + apply(simp) + apply(simp) + done + +lemma L_erase_flts: + shows "\ (L ` erase ` (set (flts rs))) = \ (L ` erase ` (set rs))" + apply(induct rs rule: flts.induct) + apply(simp_all) + apply(auto) + using L_erase_AALTs erase_fuse apply auto[1] + by (simp add: L_erase_AALTs erase_fuse) + +lemma L_erase_dB_acc: + shows "( \(L ` acc) \ ( \ (L ` erase ` (set (distinctBy rs erase acc) ) ) )) = \(L ` acc) \ \ (L ` erase ` (set rs))" + apply(induction rs arbitrary: acc) + apply simp + apply simp + by (smt (z3) SUP_absorb UN_insert sup_assoc sup_commute) + +lemma L_erase_dB: + shows " ( \ (L ` erase ` (set (distinctBy rs erase {}) ) ) ) = \ (L ` erase ` (set rs))" + by (metis L_erase_dB_acc Un_commute Union_image_empty) + +lemma L_bsimp_erase: + shows "L (erase r) = L (erase (bsimp r))" + apply(induct r) + apply(simp) + apply(simp) + apply(simp) + apply(auto simp add: Sequ_def)[1] + apply(subst L_bsimp_ASEQ[symmetric]) + apply(auto simp add: Sequ_def)[1] + apply(subst (asm) L_bsimp_ASEQ[symmetric]) + apply(auto simp add: Sequ_def)[1] + apply(simp) + apply(subst L_bsimp_AALTs[symmetric]) + defer + apply(simp) + apply(subst (2)L_erase_AALTs) + apply(subst L_erase_dB) + apply(subst L_erase_flts) + apply(auto) + apply (simp add: L_erase_AALTs) + using L_erase_AALTs by blast + + + +lemma bsimp_ASEQ0: + shows "bsimp_ASEQ bs r1 AZERO = AZERO" + apply(induct r1) + apply(auto) + done + +lemma bsimp_ASEQ1: + assumes "r1 \ AZERO" "r2 \ AZERO" "\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: + shows "bsimp_ASEQ bs (AONE bs1) r2 = fuse (bs @ bs1) r2" + apply(induct r2) + apply(auto) + done + + +lemma L_bders_simp: + shows "L (erase (bders_simp r s)) = L (erase (bders r s))" + apply(induct s arbitrary: r rule: rev_induct) + apply(simp) + apply(simp) + apply(simp add: ders_append) + apply(simp add: bders_simp_append) + apply(simp add: L_bsimp_erase[symmetric]) + by (simp add: der_correctness) + + +lemma b2: + assumes "bnullable r" + shows "bmkeps (fuse bs r) = bs @ bmkeps r" + by (simp add: assms bmkeps_retrieve bnullable_correctness erase_fuse mkeps_nullable retrieve_fuse2) + + +lemma b4: + shows "bnullable (bders_simp r s) = bnullable (bders r s)" + by (metis L_bders_simp bnullable_correctness lexer.simps(1) lexer_correct_None option.distinct(1)) + +lemma qq1: + assumes "\r \ set rs. bnullable r" + shows "bmkeps (AALTs bs (rs @ rs1)) = bmkeps (AALTs bs rs)" + using assms + apply(induct rs arbitrary: rs1 bs) + apply(simp) + apply(simp) + by (metis Nil_is_append_conv bmkeps.simps(4) neq_Nil_conv bnullable_Hdbmkeps_Hd split_list_last) + +lemma qq2: + assumes "\r \ set rs. \ bnullable r" "\r \ set rs1. bnullable r" + shows "bmkeps (AALTs bs (rs @ rs1)) = bmkeps (AALTs bs rs1)" + using assms + apply(induct rs arbitrary: rs1 bs) + apply(simp) + apply(simp) + by (metis append_assoc in_set_conv_decomp r1 r2) + +lemma qq3: + assumes "bnullable (AALTs bs (rs @ rs1))" + "bnullable (AALTs bs (rs @ rs2))" + "\bnullable (AALTs bs rs1); bnullable (AALTs bs rs2); \r\set rs. \bnullable r\ \ + bmkeps (AALTs bs rs1) = bmkeps (AALTs bs rs2)" + shows "bmkeps (AALTs bs (rs @ rs1)) = bmkeps (AALTs bs (rs @ rs2))" + using assms + apply(case_tac "\r \ set rs. bnullable r") + using qq1 apply auto[1] + by (metis UnE bnullable.simps(4) qq2 set_append) + + +lemma q3a: + assumes "\r \ set rs. bnullable r" + shows "bmkeps (AALTs bs (map (fuse bs1) rs)) = bmkeps (AALTs (bs@bs1) rs)" + using assms + apply(induct rs arbitrary: bs bs1) + apply(simp) + apply(simp) + apply(auto) + apply (metis append_assoc b2 bnullable_correctness erase_fuse bnullable_Hdbmkeps_Hd) + apply(case_tac "bnullable a") + apply (metis append.assoc b2 bnullable_correctness erase_fuse bnullable_Hdbmkeps_Hd) + apply(case_tac rs) + apply(simp) + apply(simp) + apply(auto)[1] + apply (metis bnullable_correctness erase_fuse)+ + done + +lemma qq4: + assumes "\x\set list. bnullable x" + shows "\x\set (flts list). bnullable x" + using assms + apply(induct list rule: flts.induct) + apply(auto) + by (metis UnCI bnullable_correctness erase_fuse imageI) + + +lemma qs3: + assumes "\r \ set rs. bnullable r" + shows "bmkeps (AALTs bs rs) = bmkeps (AALTs bs (flts rs))" + using assms + apply(induct rs arbitrary: bs taking: size rule: measure_induct) + apply(case_tac x) + apply(simp) + apply(simp) + apply(case_tac a) + apply(simp) + apply (simp add: r1) + apply(simp) + apply (simp add: bnullable_Hdbmkeps_Hd) + apply(simp) + apply(case_tac "flts list") + apply(simp) + apply (metis L_erase_AALTs L_erase_flts L_flat_Prf1 L_flat_Prf2 Prf_elims(1) bnullable_correctness erase.simps(4) mkeps_nullable r2) + apply(simp) + apply (simp add: r1) + prefer 3 + apply(simp) + apply (simp add: bnullable_Hdbmkeps_Hd) + prefer 2 + apply(simp) + apply(case_tac "\x\set x52. bnullable x") + apply(case_tac "list") + apply(simp) + apply (metis b2 fuse.simps(4) q3a r2) + apply(erule disjE) + apply(subst qq1) + apply(auto)[1] + apply (metis bnullable_correctness erase_fuse) + apply(simp) + apply (metis b2 fuse.simps(4) q3a r2) + apply(simp) + apply(auto)[1] + apply(subst qq1) + apply (metis bnullable_correctness erase_fuse image_eqI set_map) + apply (metis b2 fuse.simps(4) q3a r2) + apply(subst qq1) + apply (metis bnullable_correctness erase_fuse image_eqI set_map) + apply (metis b2 fuse.simps(4) q3a r2) + apply(simp) + apply(subst qq2) + apply (metis bnullable_correctness erase_fuse imageE set_map) + prefer 2 + apply(case_tac "list") + apply(simp) + apply(simp) + apply (simp add: qq4) + apply(simp) + apply(auto) + apply(case_tac list) + apply(simp) + apply(simp) + apply (simp add: bnullable_Hdbmkeps_Hd) + apply(case_tac "bnullable (ASEQ x41 x42 x43)") + apply(case_tac list) + apply(simp) + apply(simp) + apply (simp add: bnullable_Hdbmkeps_Hd) + apply(simp) + using qq4 r1 r2 by auto + +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 \ arexp \ bool" ("_ \ _" [99, 99] 99) +and + srewrite:: "arexp list \ arexp list \ bool" (" _ s\ _" [100, 100] 100) +where + bs1: "ASEQ bs AZERO r2 \ AZERO" +| bs2: "ASEQ bs r1 AZERO \ AZERO" +| bs3: "ASEQ bs1 (AONE bs2) r \ fuse (bs1@bs2) r" +| bs4: "r1 \ r2 \ ASEQ bs r1 r3 \ ASEQ bs r2 r3" +| bs5: "r3 \ r4 \ ASEQ bs r1 r3 \ ASEQ bs r1 r4" +| bs6: "AALTs bs [] \ AZERO" +| bs7: "AALTs bs [r] \ fuse bs r" +| bs10: "rs1 s\ rs2 \ AALTs bs rs1 \ AALTs bs rs2" +| ss1: "[] s\ []" +| ss2: "rs1 s\ rs2 \ (r # rs1) s\ (r # rs2)" +| ss3: "r1 \ r2 \ (r1 # rs) s\ (r2 # rs)" +| ss4: "(AZERO # rs) s\ rs" +| ss5: "(AALTs bs1 rs1 # rsb) s\ ((map (fuse bs1) rs1) @ rsb)" +| ss6: "erase a1 = erase a2 \ (rsa@[a1]@rsb@[a2]@rsc) s\ (rsa@[a1]@rsb@rsc)" + +inductive + rrewrites:: "arexp \ arexp \ bool" ("_ \* _" [100, 100] 100) +where + rs1[intro, simp]:"r \* r" +| rs2[intro]: "\r1 \* r2; r2 \ r3\ \ r1 \* r3" + +inductive + srewrites:: "arexp list \ arexp list \ bool" ("_ s\* _" [100, 100] 100) +where + sss1[intro, simp]:"rs s\* rs" +| sss2[intro]: "\rs1 s\ rs2; rs2 s\* rs3\ \ rs1 s\* rs3" + + +lemma r_in_rstar : "r1 \ r2 \ r1 \* r2" + using rrewrites.intros(1) rrewrites.intros(2) by blast + +lemma rrewrites_trans[trans]: + assumes a1: "r1 \* r2" and a2: "r2 \* r3" + shows "r1 \* r3" + using a2 a1 + apply(induct r2 r3 arbitrary: r1 rule: rrewrites.induct) + apply(auto) + done + +lemma srewrites_trans[trans]: + assumes a1: "r1 s\* r2" and a2: "r2 s\* r3" + shows "r1 s\* r3" + using a1 a2 + apply(induct r1 r2 arbitrary: r3 rule: srewrites.induct) + apply(auto) + done + + +lemma rewrite_fuse : + assumes "r1 \ r2" + shows "fuse bs r1 \ fuse bs r2" + using assms + apply(induct rule: rrewrite_srewrite.inducts(1)) + apply(auto intro: rrewrite_srewrite.intros) + apply (metis bs3 fuse_append) + by (metis bs7 fuse_append) + +lemma contextrewrites0: + "rs1 s\* rs2 \ AALTs bs rs1 \* AALTs bs rs2" + apply(induct rs1 rs2 rule: srewrites.inducts) + apply simp + using bs10 r_in_rstar rrewrites_trans by blast + +lemma contextrewrites1: + "r \* r' \ AALTs bs (r#rs) \* AALTs bs (r'#rs)" + apply(induct r r' rule: rrewrites.induct) + apply simp + using bs10 ss3 by blast + +lemma srewrite1: + shows "rs1 s\ rs2 \ (rs @ rs1) s\ (rs @ rs2)" + apply(induct rs) + apply(auto) + using ss2 by auto + +lemma srewrites1: + shows "rs1 s\* rs2 \ (rs @ rs1) s\* (rs @ rs2)" + apply(induct rs1 rs2 rule: srewrites.induct) + apply(auto) + using srewrite1 by blast + +lemma srewrite2: + shows "r1 \ r2 \ True" + and "rs1 s\ rs2 \ (rs1 @ rs) s\* (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\* rs2 \ (rs1 @ rs) s\* (rs2 @ rs)" + apply(induct rs1 rs2 arbitrary: rs rule: srewrites.induct) + apply(auto) + by (meson srewrite2(2) srewrites_trans) + +(* +lemma srewrites4: + assumes "rs3 s\* rs4" "rs1 s\* rs2" + shows "(rs1 @ rs3) s\* (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 \* r2" + shows "[r1] s\* [r2]" + using assms + + apply(induct r1 r2 rule: rrewrites.induct) + apply(auto) + by (meson srewrites.simps srewrites_trans ss3) + +lemma srewrites7: + assumes "rs3 s\* rs4" "r1 \* r2" + shows "(r1 # rs3) s\* (r2 # rs4)" + using assms + by (smt (verit, best) append_Cons append_Nil srewrites1 srewrites3 srewrites6 srewrites_trans) + + + +lemma star_seq: + assumes "r1 \* r2" + shows "ASEQ bs r1 r3 \* 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 \* r4" + shows "ASEQ bs r1 r3 \* 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 \* AZERO" + shows "ASEQ bs1 r1 r2 \* AZERO" +using assms bs1 star_seq by blast + + +lemma bsimp_aalts_simpcases: + shows "AONE bs \* bsimp (AONE bs)" + and "AZERO \* bsimp AZERO" + and "ACHAR bs c \* bsimp (ACHAR bs c)" + by (simp_all) + + +lemma trivialbsimp_srewrites: + "\\x. x \ set rs \ x \* f x \ \ rs s\* (map f rs)" + apply(induction rs) + apply simp + apply(simp) + using srewrites7 by auto + +lemma alts_simpalts: + "(\x. x \ set rs \ x \* bsimp x) \ + AALTs bs1 rs \* AALTs bs1 (map bsimp rs)" + apply(induct rs) + apply(auto)[1] + using trivialbsimp_srewrites apply auto[1] + by (simp add: contextrewrites0 srewrites7) + + +lemma bsimp_AALTs_rewrites: + shows "AALTs bs1 rs \* bsimp_AALTs bs1 rs" + by (smt (verit) bs6 bs7 bsimp_AALTs.elims rrewrites.simps) + +lemma fltsfrewrites: "rs s\* (flts rs)" + + apply(induction rs) + apply simp + apply(case_tac a) + apply(auto) + using ss4 apply blast + using srewrites7 apply force + using rs1 srewrites7 apply presburger + using srewrites7 apply force + apply (meson srewrites.simps srewrites1 ss5) + by (simp add: srewrites7) + + +lemma flts_rewrites: "AALTs bs1 rs \* AALTs bs1 (flts rs)" + by (simp add: contextrewrites0 fltsfrewrites) + + +(* delete*) +lemma threelistsappend: "rsa@a#rsb = (rsa@[a])@rsb" + apply auto + done + +lemma somewhereInside: "r \ set rs \ \rs1 rs2. rs = rs1@[r]@rs2" + using split_list by fastforce + +lemma somewhereMapInside: "f r \ f ` set rs \ \rs1 rs2 a. rs = rs1@[a]@rs2 \ f a = f r" + apply auto + by (metis split_list) + +lemma alts_dBrewrites_withFront: + "AALTs bs (rsa @ rs) \* AALTs bs (rsa @ distinctBy rs erase (erase ` set rsa))" + + apply(induction rs arbitrary: rsa) + apply simp + + apply(drule_tac x = "rsa@[a]" in meta_spec) + + apply(subst threelistsappend) + apply(rule rrewrites_trans) + apply simp + + apply(case_tac "a \ set rsa") + apply simp + apply(drule somewhereInside) + apply(erule exE)+ + apply simp + using bs10 ss6 apply auto[1] + + apply(subgoal_tac "erase ` set (rsa @ [a]) = insert (erase a) (erase ` set rsa)") + prefer 2 + + apply auto[1] + apply(case_tac "erase a \ erase `set rsa") + + apply simp + apply(subgoal_tac "AALTs bs (rsa @ a # distinctBy rs erase (insert (erase a) (erase ` set rsa))) \ + AALTs bs (rsa @ distinctBy rs erase (insert (erase a) (erase ` set rsa)))") + apply force + apply (smt (verit, ccfv_threshold) append.assoc append.left_neutral append_Cons append_Nil bs10 imageE insertCI insert_image somewhereMapInside ss6) + by simp + + + +lemma alts_dBrewrites: + shows "AALTs bs rs \* AALTs bs (distinctBy rs erase {})" + + apply(induction rs) + apply simp + apply simp + using alts_dBrewrites_withFront + by (metis append_Nil dB_single_step empty_set image_empty) + +lemma bsimp_rewrite: + shows "r \* bsimp r" +proof (induction r rule: bsimp.induct) + case (1 bs1 r1 r2) + then show "ASEQ bs1 r1 r2 \* bsimp (ASEQ bs1 r1 r2)" + apply(simp) + apply(case_tac "bsimp r1 = AZERO") + apply simp + using continuous_rewrite apply blast + apply(case_tac "\bs. bsimp r1 = AONE bs") + apply(erule exE) + apply simp + apply(subst bsimp_ASEQ2) + apply (meson rrewrites_trans rrewrite_srewrite.intros(3) rrewrites.intros(2) star_seq star_seq2) + apply (smt (verit, best) bsimp_ASEQ0 bsimp_ASEQ1 rrewrites_trans rrewrite_srewrite.intros(2) rs2 star_seq star_seq2) + done +next + case (2 bs1 rs) + then show "AALTs bs1 rs \* bsimp (AALTs bs1 rs)" + by (metis alts_dBrewrites alts_simpalts bsimp.simps(2) bsimp_AALTs_rewrites flts_rewrites rrewrites_trans) +next + case "3_1" + then show "AZERO \* bsimp AZERO" + by simp +next + case ("3_2" v) + then show "AONE v \* bsimp (AONE v)" + by simp +next + case ("3_3" v va) + then show "ACHAR v va \* bsimp (ACHAR v va)" + by simp +next + case ("3_4" v va) + then show "ASTAR v va \* bsimp (ASTAR v va)" + by simp +qed + +lemma bnullable1: +shows "r1 \ r2 \ (bnullable r1 \ bnullable r2)" + and "rs1 s\ rs2 \ ((\x \ set rs1. bnullable x) \ \x\set rs2. bnullable x)" +apply(induct rule: rrewrite_srewrite.inducts) + apply(auto) + using bnullable_fuse apply blast + apply (simp add: bnullable_fuse) + apply (meson UnCI bnullable_fuse imageI) + by (metis bnullable_correctness) + +lemma bnullable2: +shows "r1 \ r2 \ (bnullable r2 \ bnullable r1)" + and "rs1 s\ rs2 \ ((\x \ set rs2. bnullable x) \ \x\set rs1. bnullable x)" +apply(induct rule: rrewrite_srewrite.inducts) + apply(auto) + using bnullable_fuse apply blast + apply (simp add: bnullable_fuse) + using bnullable_fuse by blast + +lemma rewrite_non_nullable_strong: + assumes "r1 \ r2" + shows "bnullable r1 = bnullable r2" +using assms +apply(induction r1 r2 rule: rrewrite_srewrite.inducts(1)) +apply(auto) + apply(metis bnullable_correctness erase_fuse)+ + using bnullable1(2) apply blast + using bnullable2(2) apply blast +done + + +lemma rewritesnullable: + assumes "r1 \* r2" "bnullable r1" + shows "bnullable r2" +using assms + apply(induction r1 r2 rule: rrewrites.induct) + apply simp + using rewrite_non_nullable_strong by blast + +lemma rewrite_bmkeps_aux: + shows "r1 \ r2 \ (bnullable r1 \ bnullable r2 \ bmkeps r1 = bmkeps r2)" + and "rs1 s\ rs2 \ (\bs. (bnullable (AALTs bs rs1) \ bnullable (AALTs bs rs2) \ + bmkeps (AALTs bs rs1) = bmkeps (AALTs bs rs2)))" +proof (induct rule: rrewrite_srewrite.inducts) + case (bs1 bs r2) + then show ?case by fastforce +next + case (bs2 bs r1) + then show ?case by fastforce +next + case (bs3 bs1 bs2 r) + then show ?case by (simp add: b2) +next + case (bs4 r1 r2 bs r3) + then show ?case by simp +next + case (bs5 r3 r4 bs r1) + then show ?case by simp +next + case (bs6 bs) + then show ?case by fastforce +next + case (bs7 bs r) + then show ?case by (simp add: b2) +next + case (bs10 rs1 rs2 bs) + then show ?case + by blast +next + case ss1 + then show ?case by simp +next + case (ss2 rs1 rs2 r) + then show ?case + apply(simp) + by (metis bnullable_Hdbmkeps_Hd r1 r2) +next + case (ss3 r1 r2 rs) + then show ?case + by (metis bnullable.simps(4) bnullable_Hdbmkeps_Hd r1 rewrite_non_nullable_strong set_ConsD) +next + case (ss4 rs) + then show ?case apply(simp_all) + using bnullable.simps(1) local.ss4 r1 by blast +next + case (ss5 bs1 rs1 rsb) + then show ?case + apply(simp) + by (metis bnullable.simps(4) flts.simps(3) local.ss5 qq3 qq4 qs3) +next + case (ss6 a1 a2 bs rsa rsb rsc) + then show ?case + by (smt (verit, ccfv_threshold) append_Cons append_eq_appendI append_self_conv2 bnullable_correctness list.set_intros(1) qq3 r1) +qed + + +lemma rewrite_bmkeps: + assumes "r1 \ r2" "bnullable r1" + shows "bmkeps r1 = bmkeps r2" + using assms + by (simp add: rewrite_bmkeps_aux(1) rewrite_non_nullable_strong) + +lemma rewrites_bmkeps: + assumes "r1 \* 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 \* r2" by fact + have a3: "r2 \ r3" by fact + have a4: "bnullable r2" using a1 a2 by (simp add: rewritesnullable) + then have "bmkeps r2 = bmkeps r3" using rewrite_bmkeps a3 a4 by simp + then show "bmkeps r1 = bmkeps r3" using IH by simp +qed + +lemma to_zero_in_alt: + shows "AALT bs (ASEQ [] AZERO r) r2 \ AALT bs AZERO r2" + by (simp add: bs1 bs10 ss3) + + +lemma rewrite_fuse2: + shows "r2 \ r3 \ True" + and "rs2 s\ rs3 \ (\bs. map (fuse bs) rs2 s\* map (fuse bs) rs3)" +proof(induct rule: rrewrite_srewrite.inducts) + case ss1 + then show ?case + by simp +next + case (ss2 rs1 rs2 r) + then show ?case + using srewrites7 by force +next + case (ss3 r1 r2 rs) + then show ?case + by (simp add: r_in_rstar rewrite_fuse srewrites7) +next + case (ss4 rs) + then show ?case + by (metis fuse.simps(1) list.simps(9) rrewrite_srewrite.ss4 srewrites.simps) +next + case (ss5 bs1 rs1 rsb) + then show ?case + apply(simp) + by (metis (mono_tags, lifting) comp_def fuse_append map_eq_conv rrewrite_srewrite.ss5 srewrites.simps) +next + case (ss6 a1 a2 rsa rsb rsc) + then show ?case + apply(simp only: map_append) + by (smt (verit, ccfv_threshold) erase_fuse list.simps(8) list.simps(9) rrewrite_srewrite.ss6 srewrites.simps) +qed (auto) + + +lemma rewrites_fuse: + assumes "r1 \* r2" + shows "fuse bs r1 \* fuse bs r2" +using assms +apply(induction r1 r2 arbitrary: bs rule: rrewrites.induct) +apply(auto intro: rewrite_fuse rrewrites_trans) +done + +lemma bder_fuse_list: + shows "map (bder c \ fuse bs1) rs1 = map (fuse bs1 \ bder c) rs1" + apply(induction rs1) + apply(simp_all add: bder_fuse) + done + + +lemma rewrite_after_der: + shows "r1 \ r2 \ (bder c r1) \* (bder c r2)" + and "rs1 s\ rs2 \ map (bder c) rs1 s\* map (bder c) rs2" +proof(induction rule: rrewrite_srewrite.inducts) + case (bs1 bs r2) + then show ?case + by (simp add: continuous_rewrite) +next + case (bs2 bs r1) + then show ?case + 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) + then show ?case + apply(simp) + by (metis (no_types, lifting) bder_fuse bs10 bs7 fuse_append rrewrites.simps ss4 to_zero_in_alt) +next + case (bs4 r1 r2 bs r3) + have as: "r1 \ r2" by fact + have IH: "bder c r1 \* bder c r2" by fact + from as IH show "bder c (ASEQ bs r1 r3) \* bder c (ASEQ bs r2 r3)" + by (simp add: contextrewrites1 rewrite_bmkeps rewrite_non_nullable_strong star_seq) +next + case (bs5 r3 r4 bs r1) + have as: "r3 \ r4" by fact + have IH: "bder c r3 \* bder c r4" by fact + from as IH show "bder c (ASEQ bs r1 r3) \* 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) + then show ?case + using rrewrite_srewrite.bs6 by force +next + case (bs7 bs r) + then show ?case + by (simp add: bder_fuse r_in_rstar rrewrite_srewrite.bs7) +next + case (bs10 rs1 rs2 bs) + then show ?case + using contextrewrites0 by force +next + case ss1 + then show ?case by simp +next + case (ss2 rs1 rs2 r) + then show ?case + by (simp add: srewrites7) +next + case (ss3 r1 r2 rs) + then show ?case + by (simp add: srewrites7) +next + case (ss4 rs) + then show ?case + using rrewrite_srewrite.ss4 by fastforce +next + case (ss5 bs1 rs1 rsb) + then show ?case + apply(simp) + using bder_fuse_list map_map rrewrite_srewrite.ss5 srewrites.simps by blast +next + case (ss6 a1 a2 bs rsa rsb) + then show ?case + apply(simp only: map_append) + by (smt (verit, best) erase_bder list.simps(8) list.simps(9) local.ss6 rrewrite_srewrite.ss6 srewrites.simps) +qed + +lemma rewrites_after_der: + assumes "r1 \* r2" + shows "bder c r1 \* bder c r2" +using assms +apply(induction r1 r2 rule: rrewrites.induct) +apply(simp_all add: rewrite_after_der rrewrites_trans) +done + + +lemma central: + shows "bders r s \* bders_simp r s" +proof(induct s arbitrary: r rule: rev_induct) + case Nil + then show "bders r [] \* bders_simp r []" by simp +next + case (snoc x xs) + have IH: "\r. bders r xs \* bders_simp r xs" by fact + have "bders r (xs @ [x]) = bders (bders r xs) [x]" by (simp add: bders_append) + also have "... \* bders (bders_simp r xs) [x]" using IH + by (simp add: rewrites_after_der) + also have "... \* bders_simp (bders_simp r xs) [x]" using IH + by (simp add: bsimp_rewrite) + finally show "bders r (xs @ [x]) \* bders_simp r (xs @ [x])" + by (simp add: bders_simp_append) +qed + + + + + +lemma quasi_main: + assumes "bnullable (bders r s)" + shows "bmkeps (bders r s) = bmkeps (bders_simp r s)" +proof - + have "bders r s \* 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_main: + shows "blexer r s = blexer_simp r s" + unfolding blexer_def blexer_simp_def + using b4 quasi_main by simp + + +theorem blexersimp_correctness: + shows "lexer r s = blexer_simp r s" + using blexer_correctness main_main by simp + + + +export_code blexer_simp blexer lexer bders bders_simp in Scala module_name VerifiedLexers + + +unused_thms + + +inductive aggressive:: "arexp \ arexp \ bool" ("_ \? _" [99, 99] 99) + where + "ASEQ bs (AALTs bs1 rs) r \? AALTs (bs@bs1) (map (\r'. ASEQ [] r' r) rs) " + + + +end