# HG changeset patch # User Chengsong # Date 1646911121 0 # Node ID e072cfc2f2ee8a360b399fa93e62d67374c40c04 # Parent a7e98deebb5c0d3361d060e3bbdcd0a89050ff5e closedforms diff -r a7e98deebb5c -r e072cfc2f2ee thys2/ClosedForms.thy --- a/thys2/ClosedForms.thy Wed Mar 09 17:33:08 2022 +0000 +++ b/thys2/ClosedForms.thy Thu Mar 10 11:18:41 2022 +0000 @@ -35,7 +35,7 @@ lemma seq_closed_form: shows "rsimp (rders_simp (RSEQ r1 r2) s) = rsimp ( RALTS ( (RSEQ (rders_simp r1 s) r2) # - (map (rders r2) (vsuf s r1)) + (map (rders_simp r2) (vsuf s r1)) ) )" apply(induct s) @@ -46,6 +46,17 @@ lemma seq_closed_form_variant: shows "s \ [] \ (rders_simp (RSEQ r1 r2) s) = rsimp (RALTS ((RSEQ (rders_simp r1 s) r2) # (map (rders_simp r2) (vsuf s r1))))" - sorry + apply(induct s rule: rev_induct) + apply simp + apply(subst rders_simp_append) + apply(subst rders_simp_one_char) + apply(subst rsimp_idem[symmetric]) + apply(subst rders_simp_one_char[symmetric]) + apply(subst rders_simp_append[symmetric]) + apply(insert seq_closed_form) + apply(subgoal_tac "rsimp (rders_simp (RSEQ r1 r2) (xs @ [x])) + = rsimp (RALTS (RSEQ (rders_simp r1 (xs @ [x])) r2 # map (rders_simp r2) (vsuf (xs @ [x]) r1)))") + apply force + by presburger end \ No newline at end of file diff -r a7e98deebb5c -r e072cfc2f2ee thys2/ClosedFormsBounds.thy --- a/thys2/ClosedFormsBounds.thy Wed Mar 09 17:33:08 2022 +0000 +++ b/thys2/ClosedFormsBounds.thy Thu Mar 10 11:18:41 2022 +0000 @@ -28,6 +28,7 @@ lemma distinct_list_rexp_up_to_certain_size_bouded_by_set_enumerating_up_to_that_size: shows "\r\ set rs. (rsize r ) \ N \ sum_list (map rsize (rdistinct rs {})) \ (card (sizeNregex N))* N" + sorry @@ -88,13 +89,57 @@ sorry +lemma triangle_inequality_distinct: + shows "sum_list (map rsize (rdistinct (a # rs) ss)) \ rsize a + (sum_list (map rsize (rdistinct rs ss)))" + apply(arbitrary: ss) + apply simp + apply(case_tac "a \ ss") + apply simp + + sorry + +lemma same_regex_property_after_map: + shows "\s. P (f r2 s) \ \r \ set (map (f r2) Ss). P r" + by auto + +lemma same_property_after_distinct: + shows " \r \ set (map (f r2) Ss). P r \ \r \ set (rdistinct (map (f r2) Ss) xset). P r" + apply(induct Ss arbitrary: xset) + apply simp + by auto + +lemma same_regex_property_after_distinct: + shows "\s. P (f r2 s) \ \r \ set (rdistinct (map (f r2) Ss) xset). P r" + apply(rule same_property_after_distinct) + apply(rule same_regex_property_after_map) + by simp + +lemma map_ders_is_list_of_ders: + shows "\s. rsize (rders_simp r2 s) \ N2 \ +\r \ set (rdistinct (map (rders_simp r2) Ss) {}). rsize r \ N2" + apply(rule same_regex_property_after_distinct) + by simp + lemma seq_estimate_bounded: assumes "\s. rsize (rders_simp r1 s) \ N1" and "\s. rsize (rders_simp r2 s) \ N2" shows "Suc (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}))) \ Suc (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))" + apply(subgoal_tac " (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}))) \ + (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))") + apply force + apply(subgoal_tac " (sum_list (map rsize (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}))) \ + (rsize (RSEQ (rders_simp r1 s) r2)) + (sum_list (map rsize (rdistinct (map (rders_simp r2) (vsuf s r1)) {})) )") + prefer 2 + using triangle_inequality_distinct apply blast + apply(subgoal_tac " sum_list (map rsize (rdistinct (map (rders_simp r2) (vsuf s r1)) {})) \ N2 * card (sizeNregex N2) ") + apply(subgoal_tac "rsize (RSEQ (rders_simp r1 s) r2) \ Suc (N1 + rsize r2)") + apply linarith + apply (simp add: assms(1)) + apply(subgoal_tac "\r \ set (rdistinct (map (rders_simp r2) (vsuf s r1)) {}). rsize r \ N2") + apply (metis (no_types, opaque_lifting) assms(2) distinct_list_rexp_up_to_certain_size_bouded_by_set_enumerating_up_to_that_size ex_map_conv mult.commute) + using assms(2) map_ders_is_list_of_ders by blast - sorry lemma seq_closed_form_bounded: shows "\\s. rsize (rders_simp r1 s) \ N1 ; \s. rsize (rders_simp r2 s) \ N2\ \ @@ -145,441 +190,8 @@ - - - - - - - - - - (*Obsolete materials*) -lemma rexp_size_induct: - shows "\N r x5 a list. - \ rsize r = Suc N; r = RALTS x5; - x5 = a # list\ \\i j. rsize a = i \ rsize (RALTS list) = j \ i + j = Suc N \ i \ N \ j \ N" - apply(rule_tac x = "rsize a" in exI) - apply(rule_tac x = "rsize (RALTS list)" in exI) - apply(subgoal_tac "rsize a \ 1") - prefer 2 - using One_nat_def non_zero_size apply presburger - apply(subgoal_tac "rsize (RALTS list) \ 1 ") - prefer 2 - using size_geq1 apply blast - apply simp - done - - - - - - - - - - - - - - - - - - - - - - -lemma star_update_case1: - shows "rnullable (rders_simp r s) \ star_update c r (s # Ss) = (s @ [c]) # [c] # (star_update c r Ss)" - - by force - -lemma star_update_case2: - shows "\rnullable (rders_simp r s) \ star_update c r (s # Ss) = (s @ [c]) # (star_update c r Ss)" - by simp - -lemma bubble_break: shows "rflts [r, RZERO] = rflts [r]" - apply(case_tac r) - apply simp+ - done - -lemma rsimp_alts_idem_aux1: - shows "rsimp_ALTs (rdistinct (rflts [rsimp a]) {}) = rsimp (RALTS [a])" - by force - - - -lemma rsimp_alts_idem_aux2: - shows "rsimp a = rsimp (RALTS [a])" - apply(simp) - apply(case_tac "rsimp a") - apply simp+ - apply (metis no_alt_short_list_after_simp no_further_dB_after_simp) - by simp - -lemma rsimp_alts_idem: - shows "rsimp (rsimp_ALTs (a # as)) = rsimp (rsimp_ALTs (a # [(rsimp (rsimp_ALTs as))] ))" - apply(induct as) - apply(subgoal_tac "rsimp (rsimp_ALTs [a, rsimp (rsimp_ALTs [])]) = rsimp (rsimp_ALTs [a, RZERO])") - prefer 2 - apply simp - using bubble_break rsimp_alts_idem_aux2 apply auto[1] - apply(case_tac as) - apply(subgoal_tac "rsimp_ALTs( aa # as) = aa") - prefer 2 - apply simp - using head_one_more_simp apply fastforce - apply(subgoal_tac "rsimp_ALTs (aa # as) = RALTS (aa # as)") - prefer 2 - - using rsimp_ALTs.simps(3) apply presburger - - apply(simp only:) - apply(subgoal_tac "rsimp_ALTs (a # aa # aaa # list) = RALTS (a # aa # aaa # list)") - prefer 2 - using rsimp_ALTs.simps(3) apply presburger - apply(simp only:) - apply(subgoal_tac "rsimp_ALTs [a, rsimp (RALTS (aa # aaa # list))] = RALTS (a # [rsimp (RALTS (aa # aaa # list))])") - prefer 2 - - using rsimp_ALTs.simps(3) apply presburger - apply(simp only:) - using simp_flatten2 - apply(subgoal_tac " rsimp (RALT a (rsimp (RALTS (aa # aaa # list)))) = rsimp (RALT a ((RALTS (aa # aaa # list)))) ") - prefer 2 - - apply (metis head_one_more_simp list.simps(9) rsimp.simps(2)) - apply (simp only:) - done - - -lemma rsimp_alts_idem2: - shows "rsimp (rsimp_ALTs (a # as)) = rsimp (rsimp_ALTs ((rsimp a) # [(rsimp (rsimp_ALTs as))] ))" - using head_one_more_simp rsimp_alts_idem by auto - - -lemma evolution_step1: - shows "rsimp - (rsimp_ALTs - (rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)) # map (\s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) = - rsimp - (rsimp_ALTs - (rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)) # [(rsimp (rsimp_ALTs (map (\s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)))])) " - using rsimp_alts_idem by auto - -lemma evolution_step2: - assumes " rsimp (rsimp_ALTs (map (\s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) = - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))" - shows "rsimp - (rsimp_ALTs - (rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)) # map (\s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) = - rsimp - (rsimp_ALTs - (rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)) # [ rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))])) " - by (simp add: assms rsimp_alts_idem) - -lemma rsimp_seq_aux1: - shows "r = RONE \ r2 = RSTAR r0 \ rsimp_SEQ r r2 = r2" - apply simp - done - -lemma multiple_alts_simp_flatten: - shows "rsimp (RALT (RALT r1 r2) (rsimp_ALTs rs)) = rsimp (RALTS (r1 # r2 # rs))" - by (metis Cons_eq_appendI append_self_conv2 rsimp_ALTs.simps(2) rsimp_ALTs.simps(3) rsimp_alts_idem simp_flatten) - - -lemma evo3_main_aux1: - shows "rsimp - (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r))) - (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))) = - rsimp - (RALTS - (RSEQ (rders_simp r (a @ [x])) (RSTAR r) # - RSEQ (rders_simp r [x]) (RSTAR r) # map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))" - apply(subgoal_tac "rsimp - (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r))) - (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))) = -rsimp - (RALT (RALT (RSEQ ( (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r))) - (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))) ") - prefer 2 - apply (simp add: rsimp_idem) - apply (simp only:) - apply(subst multiple_alts_simp_flatten) - by simp - - -lemma evo3_main_nullable: - shows " -\a Ss. - \rsimp (rsimp_ALTs (map (\s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) = - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))); - rders_simp r a \ RONE; rders_simp r a \ RZERO; rnullable (rders_simp r a)\ - \ rsimp - (rsimp_ALTs - [rder x (RSEQ (rders_simp r a) (RSTAR r)), - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) = - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r (a # Ss))))" - apply(subgoal_tac "rder x (RSEQ (rders_simp r a) (RSTAR r)) - = RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ (rder x r) (RSTAR r))") - prefer 2 - apply simp - apply(simp only:) - apply(subgoal_tac "star_update x r (a # Ss) = (a @ [x]) # [x] # (star_update x r Ss)") - prefer 2 - using star_update_case1 apply presburger - apply(simp only:) - apply(subst List.list.map(2))+ - apply(subgoal_tac "rsimp - (rsimp_ALTs - [RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ (rder x r) (RSTAR r)), - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) = -rsimp - (RALTS - [RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ (rder x r) (RSTAR r)), - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))])") - prefer 2 - using rsimp_ALTs.simps(3) apply presburger - apply(simp only:) - apply(subgoal_tac " rsimp - (rsimp_ALTs - (rsimp_SEQ (rders_simp r (a @ [x])) (RSTAR r) # - rsimp_SEQ (rders_simp r [x]) (RSTAR r) # map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))) -= - rsimp - (RALTS - (rsimp_SEQ (rders_simp r (a @ [x])) (RSTAR r) # - rsimp_SEQ (rders_simp r [x]) (RSTAR r) # map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))") - - prefer 2 - using rsimp_ALTs.simps(3) apply presburger - apply (simp only:) - apply(subgoal_tac " rsimp - (RALT (RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ ( (rder x r)) (RSTAR r))) - (rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) = - rsimp - (RALT (RALT (RSEQ (rsimp (rder x (rders_simp r a))) (RSTAR r)) (RSEQ (rsimp (rder x r)) (RSTAR r))) - (rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))))") - prefer 2 - apply (simp add: rsimp_idem) - apply(simp only:) - apply(subgoal_tac " rsimp - (RALT (RALT (RSEQ (rsimp (rder x (rders_simp r a))) (RSTAR r)) (RSEQ (rsimp (rder x r)) (RSTAR r))) - (rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) = - rsimp - (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r))) - (rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))))") - prefer 2 - using rders_simp_append rders_simp_one_char rsimp_idem apply presburger - apply(simp only:) - apply(subgoal_tac " rsimp - (RALTS - (rsimp_SEQ (rders_simp r (a @ [x])) (RSTAR r) # - rsimp_SEQ (rders_simp r [x]) (RSTAR r) # map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))) = - rsimp - (RALTS - (RSEQ (rders_simp r (a @ [x])) (RSTAR r) # - RSEQ (rders_simp r [x]) (RSTAR r) # map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))") - prefer 2 - apply (smt (z3) idiot2 list.simps(9) rrexp.distinct(9) rsimp.simps(1) rsimp.simps(2) rsimp.simps(3) rsimp.simps(4) rsimp.simps(6) rsimp_idem) - apply(simp only:) - apply(subgoal_tac " rsimp - (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r))) - (rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) = - rsimp - (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r))) - ( (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) ") - prefer 2 - using rsimp_idem apply force - apply(simp only:) - using evo3_main_aux1 by blast - - -lemma evo3_main_not1: - shows " \rnullable (rders_simp r a) \ rder x (RSEQ (rders_simp r a) (RSTAR r)) = RSEQ (rder x (rders_simp r a)) (RSTAR r)" - by fastforce - - -lemma evo3_main_not2: - shows "\rnullable (rders_simp r a) \ rsimp - (rsimp_ALTs - (rder x (RSEQ (rders_simp r a) (RSTAR r)) # rs)) = rsimp - (rsimp_ALTs - ((RSEQ (rders_simp r (a @ [x])) (RSTAR r)) # rs))" - by (simp add: rders_simp_append rsimp_alts_idem2 rsimp_idem) - -lemma evo3_main_not3: - shows "rsimp - (rsimp_ALTs - (rsimp_SEQ r1 (RSTAR r) # rs)) = - rsimp (rsimp_ALTs - (RSEQ r1 (RSTAR r) # rs))" - by (metis idiot2 rrexp.distinct(9) rsimp.simps(1) rsimp.simps(3) rsimp.simps(4) rsimp.simps(6) rsimp_alts_idem rsimp_alts_idem2) - - -lemma evo3_main_notnullable: - shows "\a Ss. - \rsimp (rsimp_ALTs (map (\s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) = - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))); - rders_simp r a \ RONE; rders_simp r a \ RZERO; \rnullable (rders_simp r a)\ - \ rsimp - (rsimp_ALTs - [rder x (RSEQ (rders_simp r a) (RSTAR r)), - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) = - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r (a # Ss))))" - apply(subst star_update_case2) - apply simp - apply(subst List.list.map(2)) - apply(subst evo3_main_not2) - apply simp - apply(subst evo3_main_not3) - using rsimp_alts_idem by presburger - - -lemma evo3_aux2: - shows "rders_simp r a = RONE \ rsimp_SEQ (rders_simp (rders_simp r a) [x]) (RSTAR r) = RZERO" - by simp -lemma evo3_aux3: - shows "rsimp (rsimp_ALTs (RZERO # rs)) = rsimp (rsimp_ALTs rs)" - by (metis list.simps(8) list.simps(9) rdistinct.simps(1) rflts.simps(1) rflts.simps(2) rsimp.simps(2) rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3) rsimp_alts_idem) - -lemma evo3_aux4: - shows " rsimp - (rsimp_ALTs - [RSEQ (rder x r) (RSTAR r), - rsimp (rsimp_ALTs rs)]) = - rsimp - (rsimp_ALTs - (rsimp_SEQ (rders_simp r [x]) (RSTAR r) # rs))" - by (metis rders_simp_one_char rsimp.simps(1) rsimp.simps(6) rsimp_alts_idem rsimp_alts_idem2) - -lemma evo3_aux5: - shows "rders_simp r a \ RONE \ rders_simp r a \ RZERO \ rsimp_SEQ (rders_simp r a) (RSTAR r) = RSEQ (rders_simp r a) (RSTAR r)" - using idiot2 by blast - - -lemma evolution_step3: - shows" \a Ss. - rsimp (rsimp_ALTs (map (\s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) = - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))) \ - rsimp - (rsimp_ALTs - [rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)), - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) = - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r (a # Ss))))" - apply(case_tac "rders_simp r a = RONE") - apply(subst rsimp_seq_aux1) - apply simp - apply(subst rder.simps(6)) - apply(subgoal_tac "rnullable (rders_simp r a)") - prefer 2 - using rnullable.simps(2) apply presburger - apply(subst star_update_case1) - apply simp - - apply(subst List.list.map)+ - apply(subst rders_simp_append) - apply(subst evo3_aux2) - apply simp - apply(subst evo3_aux3) - apply(subst evo3_aux4) - apply simp - apply(case_tac "rders_simp r a = RZERO") - - apply (simp add: rsimp_alts_idem2) - apply(subgoal_tac "rders_simp r (a @ [x]) = RZERO") - prefer 2 - using rder.simps(1) rders_simp_append rders_simp_one_char rsimp.simps(3) apply presburger - using rflts.simps(2) rsimp.simps(3) rsimp_SEQ.simps(1) apply presburger - apply(subst evo3_aux5) - apply simp - apply(case_tac "rnullable (rders_simp r a) ") - using evo3_main_nullable apply blast - using evo3_main_notnullable apply blast - done - -(* -proof (prove) -goal (1 subgoal): - 1. map f (a # s) = f a # map f s -Auto solve_direct: the current goal can be solved directly with - HOL.nitpick_simp(115): map ?f (?x21.0 # ?x22.0) = ?f ?x21.0 # map ?f ?x22.0 - List.list.map(2): map ?f (?x21.0 # ?x22.0) = ?f ?x21.0 # map ?f ?x22.0 - List.list.simps(9): map ?f (?x21.0 # ?x22.0) = ?f ?x21.0 # map ?f ?x22.0 -*) -lemma starseq_list_evolution: - fixes r :: rrexp and Ss :: "char list list" and x :: char - shows "rsimp (rsimp_ALTs (map (\s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss) ) = - rsimp (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)) )" - apply(induct Ss) - apply simp - apply(subst List.list.map(2)) - apply(subst evolution_step2) - apply simp - - - sorry - - -lemma star_seqs_produce_star_seqs: - shows "rsimp (rsimp_ALTs (map (rder x \ (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) - = rsimp (rsimp_ALTs (map ( (\s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r)))) Ss))" - by (meson comp_apply) - -lemma map_der_lambda_composition: - shows "map (rder x) (map (\s. f s) Ss) = map (\s. (rder x (f s))) Ss" - by force - -lemma ralts_vs_rsimpalts: - shows "rsimp (RALTS rs) = rsimp (rsimp_ALTs rs)" - by (metis evo3_aux3 rsimp_ALTs.simps(2) rsimp_ALTs.simps(3) simp_flatten2) - - -lemma linearity_of_list_of_star_or_starseqs: - fixes r::rrexp and Ss::"char list list" and x::char - shows "\Ssa. rsimp (rder x (rsimp_ALTs (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))) = - rsimp (RALTS ( (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ssa)))" - apply(subst rder_rsimp_ALTs_commute) - apply(subst map_der_lambda_composition) - using starseq_list_evolution - apply(rule_tac x = "star_update x r Ss" in exI) - apply(subst ralts_vs_rsimpalts) - by simp - - - -(*certified correctness---does not depend on any previous sorry*) -lemma star_list_push_der: shows " \xs \ [] \ \Ss. rders_simp (RSTAR r) xs = rsimp (RALTS (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)); - xs @ [x] \ []; xs \ []\ \ - \Ss. rders_simp (RSTAR r ) (xs @ [x]) = - rsimp (RALTS (map (\s1. (rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) ) Ss) )" - apply(subgoal_tac "\Ss. rders_simp (RSTAR r) xs = rsimp (RALTS (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))") - prefer 2 - apply blast - apply(erule exE) - apply(subgoal_tac "rders_simp (RSTAR r) (xs @ [x]) = rsimp (rder x (rsimp (RALTS (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))))") - prefer 2 - using rders_simp_append - using rders_simp_one_char apply presburger - apply(rule_tac x= "Ss" in exI) - apply(subgoal_tac " rsimp (rder x (rsimp (RALTS (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)))) = - rsimp (rsimp (rder x (RALTS (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))))") - prefer 2 - using inside_simp_removal rsimp_idem apply presburger - apply(subgoal_tac "rsimp (rsimp (rder x (RALTS (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)))) = - rsimp (rsimp (RALTS (map (rder x) (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))))") - prefer 2 - using rder.simps(4) apply presburger - apply(subgoal_tac "rsimp (rsimp (RALTS (map (rder x) (map (\s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss)))) = - rsimp (rsimp (RALTS (map (\s1. (rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r)))) Ss)))") - apply (metis rsimp_idem) - by (metis map_der_lambda_composition) - - end