diff -r b794db0b79db -r b1258b7d2789 Attic/old/My.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Attic/old/My.thy Mon Jul 25 13:33:38 2011 +0000 @@ -0,0 +1,389 @@ +theory My +imports Main Infinite_Set +begin + + +definition + Seq :: "string set \ string set \ string set" ("_ ;; _" [100,100] 100) +where + "L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \ L1 \ s2 \ L2}" + +inductive_set + Star :: "string set \ string set" ("_\" [101] 102) + for L :: "string set" +where + start[intro]: "[] \ L\" +| step[intro]: "\s1 \ L; s2 \ L\\ \ s1@s2 \ L\" + +lemma lang_star_cases: + shows "L\ = {[]} \ L ;; L\" +unfolding Seq_def +by (auto) (metis Star.simps) + +lemma lang_star_cases2: + shows "L ;; L\ = L\ ;; L" +sorry + + +theorem ardens_revised: + assumes nemp: "[] \ A" + shows "(X = X ;; A \ B) \ (X = B ;; A\)" +proof + assume eq: "X = B ;; A\" + have "A\ = {[]} \ A\ ;; A" sorry + then have "B ;; A\ = B ;; ({[]} \ A\ ;; A)" unfolding Seq_def by simp + also have "\ = B \ B ;; (A\ ;; A)" unfolding Seq_def by auto + also have "\ = B \ (B ;; A\) ;; A" unfolding Seq_def + by (auto) (metis append_assoc)+ + finally show "X = X ;; A \ B" using eq by auto +next + assume "X = X ;; A \ B" + then have "B \ X" "X ;; A \ X" by auto + show "X = B ;; A\" sorry +qed + +datatype rexp = + NULL +| EMPTY +| CHAR char +| SEQ rexp rexp +| ALT rexp rexp +| STAR rexp + +fun + Sem :: "rexp \ string set" ("\_\" [0] 1000) +where + "\NULL\ = {}" + | "\EMPTY\ = {[]}" + | "\CHAR c\ = {[c]}" + | "\SEQ r1 r2\ = \r1\ ;; \r2\" + | "\ALT r1 r2\ = \r1\ \ \r2\" + | "\STAR r\ = \r\\" + +definition + folds :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" +where + "folds f z S \ SOME x. fold_graph f z S x" + +lemma folds_simp_null [simp]: + "finite rs \ x \ \folds ALT NULL rs\ \ (\r \ rs. x \ \r\)" +apply (simp add: folds_def) +apply (rule someI2_ex) +apply (erule finite_imp_fold_graph) +apply (erule fold_graph.induct) +apply (auto) +done + +lemma folds_simp_empty [simp]: + "finite rs \ x \ \folds ALT EMPTY rs\ \ (\r \ rs. x \ \r\) \ x = []" +apply (simp add: folds_def) +apply (rule someI2_ex) +apply (erule finite_imp_fold_graph) +apply (erule fold_graph.induct) +apply (auto) +done + +lemma [simp]: + shows "(x, y) \ {(x, y). P x y} \ P x y" +by simp + +definition + str_eq ("_ \_ _") +where + "x \Lang y \ (\z. x @ z \ Lang \ y @ z \ Lang)" + +definition + str_eq_rel ("\_") +where + "\Lang \ {(x, y). x \Lang y}" + +definition + final :: "string set \ string set \ bool" +where + "final X Lang \ (X \ UNIV // \Lang) \ (\s \ X. s \ Lang)" + +lemma lang_is_union_of_finals: + "Lang = \ {X. final X Lang}" +proof - + have "Lang \ \ {X. final X Lang}" + unfolding final_def + unfolding quotient_def Image_def + unfolding str_eq_rel_def + apply(simp) + apply(auto) + apply(rule_tac x="(\Lang) `` {x}" in exI) + unfolding Image_def + unfolding str_eq_rel_def + apply(auto) + unfolding str_eq_def + apply(auto) + apply(drule_tac x="[]" in spec) + apply(simp) + done + moreover + have "\{X. final X Lang} \ Lang" + unfolding final_def by auto + ultimately + show "Lang = \ {X. final X Lang}" + by blast +qed + +lemma all_rexp: + "\finite (UNIV // \Lang); X \ (UNIV // \Lang)\ \ \r. X = \r\" +sorry + +lemma final_rexp: + "\finite (UNIV // (\Lang)); final X Lang\ \ \r. X = \r\" +unfolding final_def +using all_rexp by blast + +lemma finite_f_one_to_one: + assumes "finite B" + and "\x \ B. \y. f y = x" + shows "\rs. finite rs \ (B = {f y | y . y \ rs})" +using assms +by (induct) (auto) + +lemma finite_final: + assumes "finite (UNIV // (\Lang))" + shows "finite {X. final X Lang}" +using assms +proof - + have "{X. final X Lang} \ (UNIV // (\Lang))" + unfolding final_def by auto + with assms show "finite {X. final X Lang}" + using finite_subset by auto +qed + +lemma finite_regular_aux: + fixes Lang :: "string set" + assumes "finite (UNIV // (\Lang))" + shows "\rs. Lang = \folds ALT NULL rs\" +apply(subst lang_is_union_of_finals) +using assms +apply - +apply(drule finite_final) +apply(drule_tac f="Sem" in finite_f_one_to_one) +apply(clarify) +apply(drule final_rexp[OF assms]) +apply(auto)[1] +apply(clarify) +apply(rule_tac x="rs" in exI) +apply(simp) +apply(rule set_eqI) +apply(auto) +done + +lemma finite_regular: + fixes Lang :: "string set" + assumes "finite (UNIV // (\Lang))" + shows "\r. Lang = \r\" +using assms finite_regular_aux +by auto + + + +section {* other direction *} + + +lemma inj_image_lang: + fixes f::"string \ 'a" + assumes str_inj: "\x y. f x = f y \ x \Lang y" + shows "inj_on (image f) (UNIV // (\Lang))" +proof - + { fix x y::string + assume eq_tag: "f ` {z. x \Lang z} = f ` {z. y \Lang z}" + moreover + have "{z. x \Lang z} \ {}" unfolding str_eq_def by auto + ultimately obtain a b where "x \Lang a" "y \Lang b" "f a = f b" by blast + then have "x \Lang a" "y \Lang b" "a \Lang b" using str_inj by auto + then have "x \Lang y" unfolding str_eq_def by simp + then have "{z. x \Lang z} = {z. y \Lang z}" unfolding str_eq_def by simp + } + then have "\x\UNIV // \Lang. \y\UNIV // \Lang. f ` x = f ` y \ x = y" + unfolding quotient_def Image_def str_eq_rel_def by simp + then show "inj_on (image f) (UNIV // (\Lang))" + unfolding inj_on_def by simp +qed + + +lemma finite_range_image: + assumes fin: "finite (range f)" + shows "finite ((image f) ` X)" +proof - + from fin have "finite (Pow (f ` UNIV))" by auto + moreover + have "(image f) ` X \ Pow (f ` UNIV)" by auto + ultimately show "finite ((image f) ` X)" using finite_subset by auto +qed + +definition + tag1 :: "string set \ string set \ string \ (string set \ string set)" +where + "tag1 L\<^isub>1 L\<^isub>2 \ \x. ((\L\<^isub>1) `` {x}, (\L\<^isub>2) `` {x})" + +lemma tag1_range_finite: + assumes finite1: "finite (UNIV // \L\<^isub>1)" + and finite2: "finite (UNIV // \L\<^isub>2)" + shows "finite (range (tag1 L\<^isub>1 L\<^isub>2))" +proof - + have "finite (UNIV // \L\<^isub>1 \ UNIV // \L\<^isub>2)" using finite1 finite2 by auto + moreover + have "range (tag1 L\<^isub>1 L\<^isub>2) \ (UNIV // \L\<^isub>1) \ (UNIV // \L\<^isub>2)" + unfolding tag1_def quotient_def by auto + ultimately show "finite (range (tag1 L\<^isub>1 L\<^isub>2))" + using finite_subset by blast +qed + +lemma tag1_inj: + "tag1 L\<^isub>1 L\<^isub>2 x = tag1 L\<^isub>1 L\<^isub>2 y \ x \(L\<^isub>1 \ L\<^isub>2) y" +unfolding tag1_def Image_def str_eq_rel_def str_eq_def +by auto + +lemma quot_alt_cu: + fixes L\<^isub>1 L\<^isub>2::"string set" + assumes fin1: "finite (UNIV // \L\<^isub>1)" + and fin2: "finite (UNIV // \L\<^isub>2)" + shows "finite (UNIV // \(L\<^isub>1 \ L\<^isub>2))" +proof - + have "finite (range (tag1 L\<^isub>1 L\<^isub>2))" + using fin1 fin2 tag1_range_finite by simp + then have "finite (image (tag1 L\<^isub>1 L\<^isub>2) ` (UNIV // \(L\<^isub>1 \ L\<^isub>2)))" + using finite_range_image by blast + moreover + have "\x y. tag1 L\<^isub>1 L\<^isub>2 x = tag1 L\<^isub>1 L\<^isub>2 y \ x \(L\<^isub>1 \ L\<^isub>2) y" + using tag1_inj by simp + then have "inj_on (image (tag1 L\<^isub>1 L\<^isub>2)) (UNIV // \(L\<^isub>1 \ L\<^isub>2))" + using inj_image_lang by blast + ultimately + show "finite (UNIV // \(L\<^isub>1 \ L\<^isub>2))" by (rule finite_imageD) +qed + + +section {* finite \ regular *} + +definition + transitions :: "string set \ string set \ rexp set" ("_\\_") +where + "Y \\ X \ {CHAR c | c. Y ;; {[c]} \ X}" + +definition + transitions_rexp ("_ \\ _") +where + "Y \\ X \ if [] \ X then folds ALT EMPTY (Y \\X) else folds ALT NULL (Y \\X)" + +definition + "rhs CS X \ if X = {[]} then {({[]}, EMPTY)} else {(Y, Y \\X) | Y. Y \ CS}" + +definition + "rhs_sem CS X \ \ {(Y;; \r\) | Y r . (Y, r) \ rhs CS X}" + +definition + "eqs CS \ (\X \ CS. {(X, rhs CS X)})" + +definition + "eqs_sem CS \ (\X \ CS. {(X, rhs_sem CS X)})" + +lemma [simp]: + shows "finite (Y \\ X)" +unfolding transitions_def +by auto + + +lemma defined_by_str: + assumes "s \ X" + and "X \ UNIV // (\Lang)" + shows "X = (\Lang) `` {s}" +using assms +unfolding quotient_def Image_def +unfolding str_eq_rel_def str_eq_def +by auto + +lemma every_eqclass_has_transition: + assumes has_str: "s @ [c] \ X" + and in_CS: "X \ UNIV // (\Lang)" + obtains Y where "Y \ UNIV // (\Lang)" and "Y ;; {[c]} \ X" and "s \ Y" +proof - + def Y \ "(\Lang) `` {s}" + have "Y \ UNIV // (\Lang)" + unfolding Y_def quotient_def by auto + moreover + have "X = (\Lang) `` {s @ [c]}" + using has_str in_CS defined_by_str by blast + then have "Y ;; {[c]} \ X" + unfolding Y_def Image_def Seq_def + unfolding str_eq_rel_def + by (auto) (simp add: str_eq_def) + moreover + have "s \ Y" unfolding Y_def + unfolding Image_def str_eq_rel_def str_eq_def by simp + (*moreover + have "True" by simp FIXME *) + ultimately show thesis by (blast intro: that) +qed + +lemma test: + assumes "[] \ X" + shows "[] \ \Y \\ X\" +using assms +by (simp add: transitions_rexp_def) + +lemma rhs_sem: + assumes "X \ (UNIV // (\Lang))" + shows "X \ rhs_sem (UNIV // (\Lang)) X" +apply(case_tac "X = {[]}") +apply(simp) +apply(simp add: rhs_sem_def rhs_def Seq_def) +apply(rule subsetI) +apply(case_tac "x = []") +apply(simp add: rhs_sem_def rhs_def) +apply(rule_tac x = "X" in exI) +apply(simp) +apply(rule_tac x = "X" in exI) +apply(simp add: assms) +apply(simp add: transitions_rexp_def) +oops + + +(* +fun + power :: "string \ nat \ string" (infixr "\" 100) +where + "s \ 0 = s" +| "s \ (Suc n) = s @ (s \ n)" + +definition + "Lone = {(''0'' \ n) @ (''1'' \ n) | n. True }" + +lemma + "infinite (UNIV // (\Lone))" +unfolding infinite_iff_countable_subset +apply(rule_tac x="\n. {(''0'' \ n) @ (''1'' \ i) | i. i \ {..n} }" in exI) +apply(auto) +prefer 2 +unfolding Lone_def +unfolding quotient_def +unfolding Image_def +apply(simp) +unfolding str_eq_rel_def +unfolding str_eq_def +apply(auto) +apply(rule_tac x="''0'' \ n" in exI) +apply(auto) +unfolding infinite_nat_iff_unbounded +unfolding Lone_def +*) + + + +text {* Derivatives *} + +definition + DERS :: "string \ string set \ string set" +where + "DERS s L \ {s'. s @ s' \ L}" + +lemma + shows "x \L y \ DERS x L = DERS y L" +unfolding DERS_def str_eq_def +by auto \ No newline at end of file