--- a/thys2/SizeBound6CT.thy Mon Feb 21 23:38:26 2022 +0000
+++ b/thys2/SizeBound6CT.thy Tue Mar 01 11:14:17 2022 +0000
@@ -46,7 +46,7 @@
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) )"
@@ -65,7 +65,7 @@
apply presburger
apply(drule_tac x="s1 @ [a]" in meta_spec)
sorry
-
+*)
lemma shape_of_prf_3list:
@@ -141,11 +141,6 @@
(if x \<in> acc then rdistinct xs acc
else x # (rdistinct xs ({x} \<union> acc)))"
-lemma rdistinct_idem:
- shows "rdistinct (x # (rdistinct rs {x})) {} = x # (rdistinct rs {x})"
-
- sorry
-
@@ -199,7 +194,14 @@
| "rlist_size [] = 0"
thm neq_Nil_conv
-
+lemma hand_made_def_rlist_size:
+ shows "rlist_size rs = sum_list (map rsize rs)"
+proof (induct rs)
+ case Nil show ?case by simp
+next
+ case (Cons a rs) thus ?case
+ by simp
+qed
lemma rsimp_aalts_smaller:
shows "rsize (rsimp_ALTs rs) \<le> rsize (RALTS rs)"
@@ -249,7 +251,7 @@
where
"cond_list r1 r2 s = rders_cond_list r2 (nullable_bools r1 (orderedPref s) ) (orderedSuf s)"
-thm rsimp_SEQ.simps
+
lemma rSEQ_mono:
shows "rsize (rsimp_SEQ r1 r2) \<le>rsize ( RSEQ r1 r2)"
apply auto
@@ -267,16 +269,83 @@
apply simp_all
done
+lemma ralts_cap_mono:
+ shows "rsize (RALTS rs) \<le> Suc ( sum_list (map rsize rs)) "
+ by simp
+
+lemma rflts_def_idiot:
+ shows "\<lbrakk> a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1\<rbrakk>
+ \<Longrightarrow> rflts (a # rs) = a # rflts rs"
+ apply(case_tac a)
+ apply simp_all
+ done
+
+
+lemma rflts_mono:
+ shows "sum_list (map rsize (rflts rs))\<le> sum_list (map rsize rs)"
+ apply(induct rs)
+ apply simp
+ apply(case_tac "a = RZERO")
+ apply simp
+ apply(case_tac "\<exists>rs1. a = RALTS rs1")
+ apply(erule exE)
+ apply simp
+ apply(subgoal_tac "rflts (a # rs) = a # (rflts rs)")
+ prefer 2
+ using rflts_def_idiot apply blast
+ apply simp
+ done
+
+lemma rdistinct_smaller: shows "sum_list (map rsize (rdistinct rs ss)) \<le>
+sum_list (map rsize rs )"
+ apply (induct rs arbitrary: ss)
+ apply simp
+ by (simp add: trans_le_add2)
+
+lemma rdistinct_phi_smaller: "sum_list (map rsize (rdistinct rs {})) \<le> sum_list (map rsize rs)"
+ by (simp add: rdistinct_smaller)
+
+
+lemma rsimp_alts_mono :
+ shows "\<And>x. (\<And>xa. xa \<in> set x \<Longrightarrow> rsize (rsimp xa) \<le> rsize xa) \<Longrightarrow>
+rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {})) \<le> Suc (sum_list (map rsize x))"
+ apply(subgoal_tac "rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {} ))
+ \<le> rsize (RALTS (rdistinct (rflts (map rsimp x)) {} ))")
+ prefer 2
+ using rsimp_aalts_smaller apply auto[1]
+ apply(subgoal_tac "rsize (RALTS (rdistinct (rflts (map rsimp x)) {})) \<le>Suc( sum_list (map rsize (rdistinct (rflts (map rsimp x)) {})))")
+ prefer 2
+ using ralts_cap_mono apply blast
+ apply(subgoal_tac "sum_list (map rsize (rdistinct (rflts (map rsimp x)) {})) \<le>
+ sum_list (map rsize ( (rflts (map rsimp x))))")
+ prefer 2
+ using rdistinct_smaller apply presburger
+ apply(subgoal_tac "sum_list (map rsize (rflts (map rsimp x))) \<le>
+ sum_list (map rsize (map rsimp x))")
+ prefer 2
+ using rflts_mono apply blast
+ apply(subgoal_tac "sum_list (map rsize (map rsimp x)) \<le> sum_list (map rsize x)")
+ prefer 2
+
+ apply (simp add: sum_list_mono)
+ by linarith
+
+
+
+
+
lemma rsimp_mono:
shows "rsize (rsimp r) \<le> rsize r"
+
apply(induct r)
- apply simp_all
+ apply simp_all
+
apply(subgoal_tac "rsize (rsimp_SEQ (rsimp r1) (rsimp r2)) \<le> rsize (RSEQ (rsimp r1) (rsimp r2))")
apply force
using rSEQ_mono
- apply presburger
- sorry
+ apply presburger
+ using rsimp_alts_mono by auto
lemma idiot:
shows "rsimp_SEQ RONE r = r"
@@ -284,21 +353,15 @@
apply simp_all
done
-lemma no_dup_after_simp:
- shows "RALTS rs = rsimp r \<Longrightarrow> distinct rs"
+lemma no_alt_short_list_after_simp:
+ shows "RALTS rs = rsimp r \<Longrightarrow> rsimp_ALTs rs = RALTS rs"
sorry
lemma no_further_dB_after_simp:
shows "RALTS rs = rsimp r \<Longrightarrow> rdistinct rs {} = rs"
+
sorry
-lemma 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>
@@ -552,6 +615,10 @@
RALTS (r # (cond_list r1 r2 (s @ [c]))) = RALTS (r # ((rder c r2) # (map (rder c) (cond_list r1 r2 s))))"
using suffix_plus1charn by blast
+lemma simp_flatten2:
+ shows "rsimp (RALTS (r # [RALTS rs])) = rsimp (RALTS (r # rs))"
+ sorry
+
lemma simp_flatten:
shows "rsimp (RALTS ((RALTS rsa) # rsb)) = rsimp (RALTS (rsa @ rsb))"
@@ -831,14 +898,7 @@
)"
sorry
-lemma hand_made_def_rlist_size:
- shows "rlist_size rs = sum_list (map rsize rs)"
-proof (induct rs)
- case Nil show ?case by simp
-next
- case (Cons a rs) thus ?case
- by simp
-qed
+
(*this section deals with the property of distinctBy: creates a list without duplicates*)
lemma distinct_mono:
@@ -1017,7 +1077,7 @@
"star_update c r [] = []"
|"star_update c r (s # Ss) = (if (rnullable (rders_simp r s))
then (s@[c]) # [c] # (star_update c r Ss)
- else s # (star_update c r Ss) )"
+ else (s@[c]) # (star_update c r Ss) )"
lemma star_update_case1:
shows "rnullable (rders_simp r s) \<Longrightarrow> star_update c r (s # Ss) = (s @ [c]) # [c] # (star_update c r Ss)"
@@ -1025,16 +1085,68 @@
by force
lemma star_update_case2:
- shows "\<not>rnullable (rders_simp r s) \<Longrightarrow> star_update c r (s # Ss) = s # (star_update c r Ss)"
+ shows "\<not>rnullable (rders_simp r s) \<Longrightarrow> star_update c r (s # Ss) = (s @ [c]) # (star_update c r Ss)"
+ by simp
+
+lemma bubble_break: shows "rflts [r, RZERO] = rflts [r]"
+ apply(case_tac r)
+ apply simp+
+ done
+
+lemma rsimp_alts_idem_aux1:
+ shows "rsimp_ALTs (rdistinct (rflts [rsimp a]) {}) = rsimp (RALTS [a])"
+ by force
+
+
+
+lemma rsimp_alts_idem_aux2:
+ shows "rsimp a = rsimp (RALTS [a])"
+ apply(simp)
+ apply(case_tac "rsimp a")
+ apply simp+
+ apply (metis no_alt_short_list_after_simp no_further_dB_after_simp)
by simp
lemma rsimp_alts_idem:
shows "rsimp (rsimp_ALTs (a # as)) = rsimp (rsimp_ALTs (a # [(rsimp (rsimp_ALTs as))] ))"
- sorry
+ 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))] ))"
- sorry
+ using head_one_more_simp rsimp_alts_idem by auto
+
lemma evolution_step1:
shows "rsimp
@@ -1056,6 +1168,225 @@
(rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)) # [ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))])) "
by (simp add: assms rsimp_alts_idem)
+lemma rsimp_seq_aux1:
+ shows "r = RONE \<and> r2 = RSTAR r0 \<Longrightarrow> rsimp_SEQ r r2 = r2"
+ apply simp
+ done
+
+lemma multiple_alts_simp_flatten:
+ shows "rsimp (RALT (RALT r1 r2) (rsimp_ALTs rs)) = rsimp (RALTS (r1 # r2 # rs))"
+ by (metis Cons_eq_appendI append_self_conv2 rsimp_ALTs.simps(2) rsimp_ALTs.simps(3) rsimp_alts_idem simp_flatten)
+
+
+lemma evo3_main_aux1:
+ shows "rsimp
+ (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
+ (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))) =
+ rsimp
+ (RALTS
+ (RSEQ (rders_simp r (a @ [x])) (RSTAR r) #
+ RSEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))"
+ apply(subgoal_tac "rsimp
+ (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
+ (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))) =
+rsimp
+ (RALT (RALT (RSEQ ( (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
+ (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))) ")
+ prefer 2
+ apply (simp add: rsimp_idem)
+ apply (simp only:)
+ apply(subst multiple_alts_simp_flatten)
+ by simp
+
+
+lemma evo3_main_nullable:
+ shows "
+\<And>a Ss.
+ \<lbrakk>rsimp (rsimp_ALTs (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) =
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)));
+ rders_simp r a \<noteq> RONE; rders_simp r a \<noteq> RZERO; rnullable (rders_simp r a)\<rbrakk>
+ \<Longrightarrow> rsimp
+ (rsimp_ALTs
+ [rder x (RSEQ (rders_simp r a) (RSTAR r)),
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) =
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r (a # Ss))))"
+ apply(subgoal_tac "rder x (RSEQ (rders_simp r a) (RSTAR r))
+ = RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ (rder x r) (RSTAR r))")
+ prefer 2
+ apply simp
+ apply(simp only:)
+ apply(subgoal_tac "star_update x r (a # Ss) = (a @ [x]) # [x] # (star_update x r Ss)")
+ prefer 2
+ using star_update_case1 apply presburger
+ apply(simp only:)
+ apply(subst List.list.map(2))+
+ apply(subgoal_tac "rsimp
+ (rsimp_ALTs
+ [RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ (rder x r) (RSTAR r)),
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) =
+rsimp
+ (RALTS
+ [RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ (rder x r) (RSTAR r)),
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))])")
+ prefer 2
+ using rsimp_ALTs.simps(3) apply presburger
+ apply(simp only:)
+ apply(subgoal_tac " rsimp
+ (rsimp_ALTs
+ (rsimp_SEQ (rders_simp r (a @ [x])) (RSTAR r) #
+ rsimp_SEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))
+=
+ rsimp
+ (RALTS
+ (rsimp_SEQ (rders_simp r (a @ [x])) (RSTAR r) #
+ rsimp_SEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))")
+
+ prefer 2
+ using rsimp_ALTs.simps(3) apply presburger
+ apply (simp only:)
+ apply(subgoal_tac " rsimp
+ (RALT (RALT (RSEQ (rder x (rders_simp r a)) (RSTAR r)) (RSEQ ( (rder x r)) (RSTAR r)))
+ (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) =
+ rsimp
+ (RALT (RALT (RSEQ (rsimp (rder x (rders_simp r a))) (RSTAR r)) (RSEQ (rsimp (rder x r)) (RSTAR r)))
+ (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))))")
+ prefer 2
+ apply (simp add: rsimp_idem)
+ apply(simp only:)
+ apply(subgoal_tac " rsimp
+ (RALT (RALT (RSEQ (rsimp (rder x (rders_simp r a))) (RSTAR r)) (RSEQ (rsimp (rder x r)) (RSTAR r)))
+ (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) =
+ rsimp
+ (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
+ (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))))")
+ prefer 2
+ using rders_simp_append rders_simp_one_char rsimp_idem apply presburger
+ apply(simp only:)
+ apply(subgoal_tac " rsimp
+ (RALTS
+ (rsimp_SEQ (rders_simp r (a @ [x])) (RSTAR r) #
+ rsimp_SEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))) =
+ rsimp
+ (RALTS
+ (RSEQ (rders_simp r (a @ [x])) (RSTAR r) #
+ RSEQ (rders_simp r [x]) (RSTAR r) # map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))")
+ prefer 2
+ apply (smt (z3) idiot2 list.simps(9) rrexp.distinct(9) rsimp.simps(1) rsimp.simps(2) rsimp.simps(3) rsimp.simps(4) rsimp.simps(6) rsimp_idem)
+ apply(simp only:)
+ apply(subgoal_tac " rsimp
+ (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
+ (rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) =
+ rsimp
+ (RALT (RALT (RSEQ (rsimp (rders_simp r (a @ [x]))) (RSTAR r)) (RSEQ (rders_simp r [x]) (RSTAR r)))
+ ( (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))))) ")
+ prefer 2
+ using rsimp_idem apply force
+ apply(simp only:)
+ using evo3_main_aux1 by blast
+
+
+lemma evo3_main_not1:
+ shows " \<not>rnullable (rders_simp r a) \<Longrightarrow> rder x (RSEQ (rders_simp r a) (RSTAR r)) = RSEQ (rder x (rders_simp r a)) (RSTAR r)"
+ by fastforce
+
+
+lemma evo3_main_not2:
+ shows "\<not>rnullable (rders_simp r a) \<Longrightarrow> rsimp
+ (rsimp_ALTs
+ (rder x (RSEQ (rders_simp r a) (RSTAR r)) # rs)) = rsimp
+ (rsimp_ALTs
+ ((RSEQ (rders_simp r (a @ [x])) (RSTAR r)) # rs))"
+ by (simp add: rders_simp_append rsimp_alts_idem2 rsimp_idem)
+
+lemma evo3_main_not3:
+ shows "rsimp
+ (rsimp_ALTs
+ (rsimp_SEQ r1 (RSTAR r) # rs)) =
+ rsimp (rsimp_ALTs
+ (RSEQ r1 (RSTAR r) # rs))"
+ by (metis idiot2 rrexp.distinct(9) rsimp.simps(1) rsimp.simps(3) rsimp.simps(4) rsimp.simps(6) rsimp_alts_idem rsimp_alts_idem2)
+
+
+lemma evo3_main_notnullable:
+ shows "\<And>a Ss.
+ \<lbrakk>rsimp (rsimp_ALTs (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) =
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)));
+ rders_simp r a \<noteq> RONE; rders_simp r a \<noteq> RZERO; \<not>rnullable (rders_simp r a)\<rbrakk>
+ \<Longrightarrow> rsimp
+ (rsimp_ALTs
+ [rder x (RSEQ (rders_simp r a) (RSTAR r)),
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) =
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r (a # Ss))))"
+ apply(subst star_update_case2)
+ apply simp
+ apply(subst List.list.map(2))
+ apply(subst evo3_main_not2)
+ apply simp
+ apply(subst evo3_main_not3)
+ using rsimp_alts_idem by presburger
+
+
+lemma evo3_aux2:
+ shows "rders_simp r a = RONE \<Longrightarrow> rsimp_SEQ (rders_simp (rders_simp r a) [x]) (RSTAR r) = RZERO"
+ by simp
+lemma evo3_aux3:
+ shows "rsimp (rsimp_ALTs (RZERO # rs)) = rsimp (rsimp_ALTs rs)"
+ by (metis list.simps(8) list.simps(9) rdistinct.simps(1) rflts.simps(1) rflts.simps(2) rsimp.simps(2) rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3) rsimp_alts_idem)
+
+lemma evo3_aux4:
+ shows " rsimp
+ (rsimp_ALTs
+ [RSEQ (rder x r) (RSTAR r),
+ rsimp (rsimp_ALTs rs)]) =
+ rsimp
+ (rsimp_ALTs
+ (rsimp_SEQ (rders_simp r [x]) (RSTAR r) # rs))"
+ by (metis rders_simp_one_char rsimp.simps(1) rsimp.simps(6) rsimp_alts_idem rsimp_alts_idem2)
+
+lemma evo3_aux5:
+ shows "rders_simp r a \<noteq> RONE \<and> rders_simp r a \<noteq> RZERO \<Longrightarrow> rsimp_SEQ (rders_simp r a) (RSTAR r) = RSEQ (rders_simp r a) (RSTAR r)"
+ using idiot2 by blast
+
+
+lemma evolution_step3:
+ shows" \<And>a Ss.
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rder x (rsimp_SEQ (rders_simp r s1) (RSTAR r))) Ss)) =
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss))) \<Longrightarrow>
+ rsimp
+ (rsimp_ALTs
+ [rder x (rsimp_SEQ (rders_simp r a) (RSTAR r)),
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r Ss)))]) =
+ rsimp (rsimp_ALTs (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) (star_update x r (a # Ss))))"
+ apply(case_tac "rders_simp r a = RONE")
+ apply(subst rsimp_seq_aux1)
+ apply simp
+ apply(subst rder.simps(6))
+ apply(subgoal_tac "rnullable (rders_simp r a)")
+ prefer 2
+ using rnullable.simps(2) apply presburger
+ apply(subst star_update_case1)
+ apply simp
+
+ apply(subst List.list.map)+
+ apply(subst rders_simp_append)
+ apply(subst evo3_aux2)
+ apply simp
+ apply(subst evo3_aux3)
+ apply(subst evo3_aux4)
+ apply simp
+ apply(case_tac "rders_simp r a = RZERO")
+
+ apply (simp add: rsimp_alts_idem2)
+ apply(subgoal_tac "rders_simp r (a @ [x]) = RZERO")
+ prefer 2
+ using rder.simps(1) rders_simp_append rders_simp_one_char rsimp.simps(3) apply presburger
+ using rflts.simps(2) rsimp.simps(3) rsimp_SEQ.simps(1) apply presburger
+ apply(subst evo3_aux5)
+ apply simp
+ apply(case_tac "rnullable (rders_simp r a) ")
+ using evo3_main_nullable apply blast
+ using evo3_main_notnullable apply blast
+ done
(*
proof (prove)
@@ -1075,11 +1406,8 @@
apply(subst List.list.map(2))
apply(subst evolution_step2)
apply simp
- apply(case_tac "rnullable (rders_simp r a)")
- apply(subst star_update_case1)
- apply simp
- apply(subst List.list.map)+
- sledgehammer
+
+
sorry
@@ -1094,7 +1422,8 @@
lemma ralts_vs_rsimpalts:
shows "rsimp (RALTS rs) = rsimp (rsimp_ALTs rs)"
- sorry
+ 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
@@ -1156,14 +1485,119 @@
using star_list_push_der apply presburger
+ by (metis ralts_vs_rsimpalts starseq_list_evolution)
+
+
+lemma starder_is_a_list:
+ shows " \<exists>Ss. rders_simp (RSTAR r) s = rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r s1) (RSTAR r)) Ss))) \<or> rders_simp (RSTAR r) s = RSTAR r"
+ apply(case_tac s)
+ prefer 2
+ apply (metis neq_Nil_conv starder_is_a_list_of_stars_or_starseqs)
+ apply simp
+ done
+
+
+(** start about bounds here**)
+
+
+lemma list_simp_size:
+ shows "rlist_size (map rsimp rs) \<le> rlist_size rs"
+ apply(induct rs)
+ apply simp
+ apply simp
+ apply (subgoal_tac "rsize (rsimp a) \<le> rsize a")
+ prefer 2
+ using rsimp_mono apply fastforce
+ using add_le_mono by presburger
+
+lemma inside_list_simp_inside_list:
+ shows "r \<in> set rs \<Longrightarrow> rsimp r \<in> set (map rsimp rs)"
+ apply (induct rs)
+ apply simp
+ apply auto
+ done
+
+
+lemma rsize_star_seq_list:
+ shows "(\<forall>s. rsize (rders_simp r0 s) < N0 ) \<Longrightarrow> \<exists>N3.\<forall>Ss.
+rlist_size (rdistinct (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss) {}) < N3"
sorry
+lemma rdistinct_bound_by_no_simp:
+ shows "
+
+ rlist_size (rdistinct (map rsimp rs) (set (map rsimp ss)))
+ \<le> (rlist_size (rdistinct rs (set ss)))
+"
+ apply(induct rs arbitrary: ss)
+ apply simp
+ apply(case_tac "a \<in> set ss")
+ apply(subgoal_tac "rsimp a \<in> set (map rsimp ss)")
+ prefer 2
+ using inside_list_simp_inside_list apply blast
+
+ apply simp
+ apply simp
+ by (metis List.set_insert add_le_mono image_insert insert_absorb rsimp_mono trans_le_add2)
+
+
+lemma starder_closed_form_bound_aux1:
+ shows
+"\<forall>Ss. rsize (rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) \<le>
+ Suc (rlist_size ( (rdistinct ( ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss))) {}))) "
+
+ sorry
+
+lemma starder_closed_form_bound:
+ shows "(\<forall>s. rsize (rders_simp r0 s) < N0 ) \<Longrightarrow> \<exists>N3.\<forall>Ss.
+rsize(rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) < N3"
+ apply(subgoal_tac " \<exists>N3.\<forall>Ss.
+rlist_size (rdistinct (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss) {}) < N3")
+ prefer 2
+
+ using rsize_star_seq_list apply auto[1]
+ apply(erule exE)
+ apply(rule_tac x = "Suc N3" in exI)
+ apply(subgoal_tac "\<forall>Ss. rsize (rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) \<le>
+ Suc (rlist_size ( (rdistinct ( ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss))) {})))")
+ prefer 2
+ using starder_closed_form_bound_aux1 apply blast
+ by (meson less_trans_Suc linorder_not_le not_less_eq)
+
+
+thm starder_closed_form_bound_aux1
+
+(*
+ "ralts_vs_rsimpalts", , and "starder_closed_form_bound_aux1", which could be due to a bug in Sledgehammer or to inconsistent axioms (including "sorry"s)
+*)
+
+lemma starder_size_bound:
+ shows "(\<forall>s. rsize (rders_simp r0 s) < N0 ) \<Longrightarrow> \<exists>N3.\<forall>Ss.
+rsize(rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) < N3 \<and>
+rsize (RSTAR r0) < N3"
+ apply(subgoal_tac " \<exists>N3.\<forall>Ss.
+rsize(rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) < N3")
+ prefer 2
+ using starder_closed_form_bound apply blast
+ apply(erule exE)
+ apply(rule_tac x = "max N3 (Suc (rsize (RSTAR r0)))" in exI)
+ using less_max_iff_disj by blast
+
+
+
+
lemma finite_star:
shows "(\<forall>s. rsize (rders_simp r0 s) < N0 )
\<Longrightarrow> \<exists>N3. \<forall>s.(rsize (rders_simp (RSTAR r0) s)) < N3"
-
- sorry
+ apply(subgoal_tac " \<exists>N3. \<forall>Ss.
+rsize(rsimp (RALTS ( (map (\<lambda>s1. rsimp_SEQ (rders_simp r0 s1) (RSTAR r0)) Ss)))) < N3 \<and>
+rsize (RSTAR r0) < N3")
+ prefer 2
+ using starder_size_bound apply blast
+ apply(erule exE)
+ apply(rule_tac x = N3 in exI)
+ by (metis starder_is_a_list)
lemma rderssimp_zero: