diff -r 232aa2f19a75 -r ec5e4fe4cc70 thys2/Exercises.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys2/Exercises.thy Sun Oct 10 18:35:21 2021 +0100 @@ -0,0 +1,253 @@ +theory Exercises + imports Spec "~~/src/HOL/Library/Infinite_Set" +begin + +section {* Some Fun Facts *} + +fun + zeroable :: "rexp \ bool" +where + "zeroable (ZERO) \ True" +| "zeroable (ONE) \ False" +| "zeroable (CH c) \ False" +| "zeroable (ALT r1 r2) \ zeroable r1 \ zeroable r2" +| "zeroable (SEQ r1 r2) \ zeroable r1 \ zeroable r2" +| "zeroable (STAR r) \ False" + +lemma zeroable_correctness: + shows "zeroable r \ L r = {}" +by(induct r)(auto simp add: Sequ_def) + + +fun + atmostempty :: "rexp \ bool" +where + "atmostempty (ZERO) \ True" +| "atmostempty (ONE) \ True" +| "atmostempty (CH c) \ False" +| "atmostempty (ALT r1 r2) \ atmostempty r1 \ atmostempty r2" +| "atmostempty (SEQ r1 r2) \ + zeroable r1 \ zeroable r2 \ (atmostempty r1 \ atmostempty r2)" +| "atmostempty (STAR r) = atmostempty r" + + + +fun + somechars :: "rexp \ bool" +where + "somechars (ZERO) \ False" +| "somechars (ONE) \ False" +| "somechars (CH c) \ True" +| "somechars (ALT r1 r2) \ somechars r1 \ somechars r2" +| "somechars (SEQ r1 r2) \ + (\ zeroable r1 \ somechars r2) \ (\ zeroable r2 \ somechars r1) \ + (somechars r1 \ nullable r2) \ (somechars r2 \ nullable r1)" +| "somechars (STAR r) \ somechars r" + +lemma somechars_correctness: + shows "somechars r \ (\s. s \ [] \ s \ L r)" +apply(induct r) +apply(simp_all add: zeroable_correctness nullable_correctness Sequ_def) +using Nil_is_append_conv apply blast +apply blast + apply(auto) + by (metis Star_decomp hd_Cons_tl list.distinct(1)) + +lemma atmostempty_correctness_aux: + shows "atmostempty r \ \ somechars r" +apply(induct r) +apply(simp_all) +apply(auto simp add: zeroable_correctness nullable_correctness somechars_correctness) +done + +lemma atmostempty_correctness: + shows "atmostempty r \ L r \ {[]}" +by(auto simp add: atmostempty_correctness_aux somechars_correctness) + +fun + leastsinglechar :: "rexp \ bool" +where + "leastsinglechar (ZERO) \ False" +| "leastsinglechar (ONE) \ False" +| "leastsinglechar (CH c) \ True" +| "leastsinglechar (ALT r1 r2) \ leastsinglechar r1 \ leastsinglechar r2" +| "leastsinglechar (SEQ r1 r2) \ + (if (zeroable r1 \ zeroable r2) then False + else ((nullable r1 \ leastsinglechar r2) \ (nullable r2 \ leastsinglechar r1)))" +| "leastsinglechar (STAR r) \ leastsinglechar r" + +lemma leastsinglechar_correctness: + "leastsinglechar r \ (\c. [c] \ L r)" + apply(induct r) + apply(simp) + apply(simp) + apply(simp) + prefer 2 + apply(simp) + apply(blast) + prefer 2 + apply(simp) + using Star.step Star_decomp apply fastforce + apply(simp add: Sequ_def zeroable_correctness nullable_correctness) + by (metis append_Nil append_is_Nil_conv butlast_append butlast_snoc) + +fun + infinitestrings :: "rexp \ bool" +where + "infinitestrings (ZERO) = False" +| "infinitestrings (ONE) = False" +| "infinitestrings (CH c) = False" +| "infinitestrings (ALT r1 r2) = (infinitestrings r1 \ infinitestrings r2)" +| "infinitestrings (SEQ r1 r2) \ + (\ zeroable r1 \ infinitestrings r2) \ (\ zeroable r2 \ infinitestrings r1)" +| "infinitestrings (STAR r) = (\ atmostempty r)" + + + + + +lemma Star_atmostempty: + assumes "A \ {[]}" + shows "A\ \ {[]}" + using assms + using Star_decomp concat_eq_Nil_conv empty_iff insert_iff subsetI subset_singletonD + apply(auto) +proof - + fix x :: "char list" + assume a1: "x \ A\" + assume "\c x A. c # x \ A\ \ \s1 s2. x = s1 @ s2 \ c # s1 \ A \ s2 \ A\" + then have f2: "\cs C c. \csa. c # csa \ C \ c # cs \ C\" + by auto + obtain cc :: "char list \ char" and ccs :: "char list \ char list" where + "\cs. cs = [] \ cc cs # ccs cs = cs" + by (metis (no_types) list.exhaust) + then show "x = []" + using f2 a1 by (metis assms empty_iff insert_iff list.distinct(1) subset_singletonD) +qed + + +lemma Star_empty_string_finite: + shows "finite ({[]}\)" +using Star_atmostempty infinite_super by auto + +lemma Star_empty_finite: + shows "finite ({}\)" +using Star_atmostempty infinite_super by auto + +lemma Star_concat_replicate: + assumes "s \ A" + shows "concat (replicate n s) \ A\" +using assms +by (induct n) (auto) + + +lemma concat_replicate_inj: + assumes "concat (replicate n s) = concat (replicate m s)" "s \ []" + shows "n = m" +using assms +apply(induct n arbitrary: m) +apply(auto)[1] +apply(auto) +apply(case_tac m) +apply(clarify) +apply(simp only: replicate.simps concat.simps) +apply blast +by simp + +lemma A0: + assumes "finite (A ;; B)" "B \ {}" + shows "finite A" +apply(subgoal_tac "\s. s \ B") +apply(erule exE) +apply(subgoal_tac "finite {s1 @ s |s1. s1 \ A}") +apply(rule_tac f="\s1. s1 @ s" in finite_imageD) +apply(simp add: image_def) +apply(smt Collect_cong) +apply(simp add: inj_on_def) +apply(rule_tac B="A ;; B" in finite_subset) +apply(auto simp add: Sequ_def)[1] +apply(rule assms(1)) +using assms(2) by auto + +lemma A1: + assumes "finite (A ;; B)" "A \ {}" + shows "finite B" +apply(subgoal_tac "\s. s \ A") +apply(erule exE) +apply(subgoal_tac "finite {s @ s1 |s1. s1 \ B}") +apply(rule_tac f="\s1. s @ s1" in finite_imageD) +apply(simp add: image_def) +apply(smt Collect_cong) +apply(simp add: inj_on_def) +apply(rule_tac B="A ;; B" in finite_subset) +apply(auto simp add: Sequ_def)[1] +apply(rule assms(1)) +using assms(2) by auto + +lemma Sequ_Prod_finite: + assumes "A \ {}" "B \ {}" + shows "finite (A ;; B) \ (finite (A \ B))" +apply(rule iffI) +apply(rule finite_cartesian_product) +apply(erule A0) +apply(rule assms(2)) +apply(erule A1) +apply(rule assms(1)) +apply(simp add: Sequ_def) +apply(rule finite_image_set2) +apply(drule finite_cartesian_productD1) +apply(rule assms(2)) +apply(simp) +apply(drule finite_cartesian_productD2) +apply(rule assms(1)) +apply(simp) +done + + +lemma Star_non_empty_string_infinite: + assumes "s \ A" " s \ []" + shows "infinite (A\)" +proof - + have "inj (\n. concat (replicate n s))" + using assms(2) concat_replicate_inj + by(auto simp add: inj_on_def) + moreover + have "infinite (UNIV::nat set)" by simp + ultimately + have "infinite ((\n. concat (replicate n s)) ` UNIV)" + by (simp add: range_inj_infinite) + moreover + have "((\n. concat (replicate n s)) ` UNIV) \ (A\)" + using Star_concat_replicate assms(1) by auto + ultimately show "infinite (A\)" + using infinite_super by auto +qed + +lemma infinitestrings_correctness: + shows "infinitestrings r \ infinite (L r)" +apply(induct r) +apply(simp_all) +apply(simp add: zeroable_correctness) +apply(rule iffI) +apply(erule disjE) +apply(subst Sequ_Prod_finite) +apply(auto)[2] +using finite_cartesian_productD2 apply blast +apply(subst Sequ_Prod_finite) +apply(auto)[2] +using finite_cartesian_productD1 apply blast +apply(subgoal_tac "L r1 \ {} \ L r2 \ {}") +prefer 2 +apply(auto simp add: Sequ_def)[1] +apply(subst (asm) Sequ_Prod_finite) +apply(auto)[2] +apply(auto)[1] +apply(simp add: atmostempty_correctness) +apply(rule iffI) +apply (metis Star_empty_finite Star_empty_string_finite subset_singletonD) +using Star_non_empty_string_infinite apply blast +done + +unused_thms + +end \ No newline at end of file