# HG changeset patch # User urbanc # Date 1288098082 0 # Node ID f8a6ea83f7b60195b589bfbe883e9d0d755affed # Parent a761b8ac8488710094702cf4563e5a0fd3c8526b more experiments diff -r a761b8ac8488 -r f8a6ea83f7b6 MyhillNerode.thy --- a/MyhillNerode.thy Sun Oct 24 01:17:44 2010 +0000 +++ b/MyhillNerode.thy Tue Oct 26 13:01:22 2010 +0000 @@ -1312,42 +1312,120 @@ text {* tests by Christian *} +(* compatibility with stable Isabelle *) lemma temp_set_ext[no_atp]: "(\x. (x:A) = (x:B)) \ (A = B)" -by(auto intro:set_eqI) + by(auto intro:set_eqI) + +abbreviation + str_eq ("_ \_ _") +where + "x \Lang y \ (\z. x @ z \ Lang \ y @ z \ Lang)" + +abbreviation + lang_eq ("\_") +where + "\Lang \ (\(x, y). x \Lang y)" -fun - MNRel -where - "MNRel Lang (x, y) = (x \Lang\ y)" +lemma lang_eq_is_equiv: + "equiv UNIV (\Lang)" +unfolding equiv_def +unfolding refl_on_def sym_def trans_def +by (simp add: mem_def equiv_str_def) + +text {* + equivalence classes of x can be written + + (\Lang)``{x} + + the language quotient can be written + + UNIV // (\Lang) +*} + +lemma + "(\Lang)``{x} = \x\Lang" +unfolding equiv_class_def equiv_str_def Image_def mem_def +by (simp) lemma - "equiv UNIV (MNRel Lang)" -unfolding equiv_def -unfolding refl_on_def sym_def trans_def -apply(auto simp add: mem_def equiv_str_def) + "UNIV // (\Lang) = UNIV Quo Lang" +unfolding quotient_def quot_def +unfolding equiv_class_def equiv_str_def +unfolding Image_def mem_def +by auto + +lemma + "{} \ UNIV // (\Lang)" +unfolding quotient_def +unfolding Image_def +apply(auto) +apply(rule_tac x="x" in exI) +unfolding mem_def +apply(simp) +done + +definition + transitions :: "string set \ string set \ rexp set" ("_\\_") +where + "X \\ Y \ {CHAR c | c. X ; {[c]} \ Y}" + +definition + "rhs CS X \ { (Y, X \\Y) | Y. Y \ CS }" + +definition + "rhs_sem CS X \ { (Y; L (folds ALT NULL (X \\Y))) | Y. Y \ CS }" + +definition + "eqs CS \ (\X \ CS. {(X, rhs CS X)})" + +lemma finite_rhs: + assumes fin: "finite CS" + shows "finite (rhs CS X)" + using fin unfolding rhs_def by (auto) + +lemma finite_eqs: + assumes fin: "finite CS" + shows "finite (eqs CS)" +unfolding eqs_def +apply(rule finite_UN_I) +apply(rule fin) +apply(simp add: eq_def) done lemma - "(MNRel Lang)``{x} = \x\Lang" -unfolding equiv_class_def Image_def mem_def -by (simp) + shows "X ; L (folds ALT NULL X \\ Y) \ Y" +by (auto simp add: transitions_def lang_seq_def fold_alt_null_eqs) + +lemma + assumes a: "X \ (UNIV // (\Lang))" + shows "X \ \ (rhs_sem (UNIV // (\Lang)) X)" +unfolding rhs_sem_def +apply (auto simp add: transitions_def lang_seq_def fold_alt_null_eqs) +apply(rule_tac x="X" in exI) +apply(simp) +apply(simp add: quotient_def Image_def) +apply(subst (4) mem_def) +apply(simp) +using a +apply(simp add: quotient_def Image_def) +apply(subst (asm) (2) mem_def) +apply(simp) +apply(rule_tac x="X" in exI) +apply(simp) +apply(erule exE) +apply(simp) +apply(rule temp_set_ext) +apply(simp) +apply(rule iffI) +apply(rule_tac x="x" in exI) +apply(simp) +oops lemma tt: "\A = B; C \ B\ \ C \ A" by simp lemma - "UNIV//(MNRel (L1 \ L2)) \ (UNIV//(MNRel L1)) \ (UNIV//(MNRel L2))" -unfolding quotient_def -unfolding UNION_eq_Union_image -apply(rule tt) -apply(rule Un_Union_image[symmetric]) -apply(simp) -apply(rule UN_mono) -apply(simp) -apply(simp) -unfolding Image_def -unfolding mem_def -unfolding MNRel.simps -apply(simp) + "UNIV//(\(L1 \ L2)) \ (UNIV//(\L1)) \ (UNIV//(\L2))" +unfolding quotient_def Image_def oops @@ -1358,13 +1436,13 @@ where "QUOT Lang \ (\x. {\x\Lang})" + lemma test_01: "Lang \ (\ QUOT Lang)" unfolding QUOT_def equiv_class_def equiv_str_def by (blast) - -(* by chunhan *) +text{* by chunhan *} lemma quot_lambda: "QUOT {[]} = {{[]}, UNIV - {[]}}" proof show "QUOT {[]} \ {{[]}, UNIV - {[]}}"