diff -r c730d018ebfa -r f9cdc295ccf7 thys3/GeneralRegexBound.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/GeneralRegexBound.thy Thu Apr 28 15:56:22 2022 +0100 @@ -0,0 +1,212 @@ +theory GeneralRegexBound + imports "BasicIdentities" +begin + +lemma size_geq1: + shows "rsize r \ 1" + by (induct r) auto + +definition RSEQ_set where + "RSEQ_set A n \ {RSEQ r1 r2 | r1 r2. r1 \ A \ r2 \ A \ rsize r1 + rsize r2 \ n}" + +definition RSEQ_set_cartesian where + "RSEQ_set_cartesian A = {RSEQ r1 r2 | r1 r2. r1 \ A \ r2 \ A}" + +definition RALT_set where + "RALT_set A n \ {RALTS rs | rs. set rs \ A \ rsizes rs \ n}" + +definition RALTs_set where + "RALTs_set A n \ {RALTS rs | rs. \r \ set rs. r \ A \ rsizes rs \ n}" + +definition + "sizeNregex N \ {r. rsize r \ N}" + + +lemma sizenregex_induct1: + "sizeNregex (Suc n) = (({RZERO, RONE} \ {RCHAR c| c. True}) + \ (RSTAR ` sizeNregex n) + \ (RSEQ_set (sizeNregex n) n) + \ (RALTs_set (sizeNregex n) n))" + apply(auto) + apply(case_tac x) + apply(auto simp add: RSEQ_set_def) + using sizeNregex_def apply force + using sizeNregex_def apply auto[1] + apply (simp add: sizeNregex_def) + apply (simp add: sizeNregex_def) + apply (simp add: RALTs_set_def) + apply (metis imageI list.set_map member_le_sum_list order_trans) + apply (simp add: sizeNregex_def) + apply (simp add: sizeNregex_def) + apply (simp add: sizeNregex_def) + using sizeNregex_def apply force + apply (simp add: sizeNregex_def) + apply (simp add: sizeNregex_def) + apply (simp add: RALTs_set_def) + apply(simp add: sizeNregex_def) + apply(auto) + using ex_in_conv by fastforce + +lemma s4: + "RSEQ_set A n \ RSEQ_set_cartesian A" + using RSEQ_set_cartesian_def RSEQ_set_def by fastforce + +lemma s5: + assumes "finite A" + shows "finite (RSEQ_set_cartesian A)" + using assms + apply(subgoal_tac "RSEQ_set_cartesian A = (\(x1, x2). RSEQ x1 x2) ` (A \ A)") + apply simp + unfolding RSEQ_set_cartesian_def + apply(auto) + done + + +definition RALTs_set_length + where + "RALTs_set_length A n l \ {RALTS rs | rs. \r \ set rs. r \ A \ rsizes rs \ n \ length rs \ l}" + + +definition RALTs_set_length2 + where + "RALTs_set_length2 A l \ {RALTS rs | rs. \r \ set rs. r \ A \ length rs \ l}" + +definition set_length2 + where + "set_length2 A l \ {rs. \r \ set rs. r \ A \ length rs \ l}" + + +lemma r000: + shows "RALTs_set_length A n l \ RALTs_set_length2 A l" + apply(auto simp add: RALTs_set_length2_def RALTs_set_length_def) + done + + +lemma r02: + shows "set_length2 A 0 \ {[]}" + apply(auto simp add: set_length2_def) + apply(case_tac x) + apply(auto) + done + +lemma r03: + shows "set_length2 A (Suc n) \ + {[]} \ (\(h, t). h # t) ` (A \ (set_length2 A n))" + apply(auto simp add: set_length2_def) + apply(case_tac x) + apply(auto) + done + +lemma r1: + assumes "finite A" + shows "finite (set_length2 A n)" + using assms + apply(induct n) + apply(rule finite_subset) + apply(rule r02) + apply(simp) + apply(rule finite_subset) + apply(rule r03) + apply(simp) + done + +lemma size_sum_more_than_len: + shows "rsizes rs \ length rs" + apply(induct rs) + apply simp + apply simp + apply(subgoal_tac "rsize a \ 1") + apply linarith + using size_geq1 by auto + + +lemma sum_list_len: + shows "rsizes rs \ n \ length rs \ n" + by (meson order.trans size_sum_more_than_len) + + +lemma t2: + shows "RALTs_set A n \ RALTs_set_length A n n" + unfolding RALTs_set_length_def RALTs_set_def + apply(auto) + using sum_list_len by blast + +lemma s8_aux: + assumes "finite A" + shows "finite (RALTs_set_length A n n)" +proof - + have "finite A" by fact + then have "finite (set_length2 A n)" + by (simp add: r1) + moreover have "(RALTS ` (set_length2 A n)) = RALTs_set_length2 A n" + unfolding RALTs_set_length2_def set_length2_def + by (auto) + ultimately have "finite (RALTs_set_length2 A n)" + by (metis finite_imageI) + then show ?thesis + by (metis infinite_super r000) +qed + +lemma char_finite: + shows "finite {RCHAR c |c. True}" + apply simp + apply(subgoal_tac "finite (RCHAR ` (UNIV::char set))") + prefer 2 + apply simp + by (simp add: full_SetCompr_eq) + + +lemma finite_size_n: + shows "finite (sizeNregex n)" + apply(induct n) + apply(simp add: sizeNregex_def) + apply (metis (mono_tags, lifting) not_finite_existsD not_one_le_zero size_geq1) + apply(subst sizenregex_induct1) + apply(simp only: finite_Un) + apply(rule conjI)+ + apply(simp) + + using char_finite apply blast + apply(simp) + apply(rule finite_subset) + apply(rule s4) + apply(rule s5) + apply(simp) + apply(rule finite_subset) + apply(rule t2) + apply(rule s8_aux) + apply(simp) + done + +lemma three_easy_cases0: + shows "rsize (rders_simp RZERO s) \ Suc 0" + apply(induct s) + apply simp + apply simp + done + + +lemma three_easy_cases1: + shows "rsize (rders_simp RONE s) \ Suc 0" + apply(induct s) + apply simp + apply simp + using three_easy_cases0 by auto + + +lemma three_easy_casesC: + shows "rsize (rders_simp (RCHAR c) s) \ Suc 0" + apply(induct s) + apply simp + apply simp + apply(case_tac " a = c") + using three_easy_cases1 apply blast + apply simp + using three_easy_cases0 by force + + +unused_thms + + +end +