# HG changeset patch # User urbanc # Date 1314965940 0 # Node ID 77d9ed8f5c8445ff0e210a037dfa41078ae00341 # Parent 5f7b7ad498dd756a1dde554f84c4752c9bdc5d8c added example about non-regularity diff -r 5f7b7ad498dd -r 77d9ed8f5c84 Closures.thy --- a/Closures.thy Fri Sep 02 09:53:56 2011 +0000 +++ b/Closures.thy Fri Sep 02 12:19:00 2011 +0000 @@ -164,7 +164,6 @@ then show "regular (Deriv_lang B A)" by auto qed - subsection {* Finite and co-finite sets are regular *} lemma singleton_regular: @@ -213,7 +212,7 @@ subsection {* Non-regularity for languages *} lemma continuation_lemma: - fixes A B::"('a::finite list) set" + fixes A B::"'a::finite lang" assumes reg: "regular A" and inf: "infinite B" shows "\x \ B. \y \ B. x \ y \ x \A y" @@ -244,31 +243,57 @@ qed +subsubsection {* The language @{text "a ^ n b ^ n"} is not regular *} -(* tests *) -definition - "quot A B \ {x. \y \ B. x @ y \ A}" +abbreviation + replicate_rev ("_ ^^^ _" [100, 100] 100) +where + "a ^^^ n \ replicate n a" definition - "quot1 A B \ {x. \y \ B. y @ x \ A}" + "length_test s a b \ length (filter (op= a) s) = length (filter (op= b) s)" + +lemma [simp]: + assumes "i \ j" "a \ b" + shows "length_test ((a ^^^ i) @ (b ^^^ j)) a b \ length_test ((a ^^^ n) @ (b ^^^ n)) a b" +using assms unfolding length_test_def by auto -lemma - "quot1 A B \ Deriv_lang B A" -unfolding quot1_def Derivs_def -apply(auto) -done +lemma [simp]: + assumes "a \ b" + shows "filter (op= a) ((a ^^^ i) @ (b ^^^ j)) = a ^^^ i" + and "filter (op= b) ((a ^^^ i) @ (b ^^^ j)) = b ^^^ j" +using assms by simp_all + +lemma length_test: + "x = y \ length_test x a b = length_test y a b" +by simp -lemma - "rev ` quot1 A B \ quot (rev ` A) (rev ` B)" -unfolding quot_def quot1_def -apply(auto) -unfolding image_def -apply(auto) -apply(rule_tac x="y" in bexI) -apply(rule_tac x="y @ xa" in bexI) -apply(auto) -done +lemma an_bn_not_regular: + shows "\ regular (\n. {CHR ''a'' ^^^ n @ CHR ''b'' ^^^ n})" +proof + def A\"\n. {CHR ''a'' ^^^ n @ CHR ''b'' ^^^ n}" + def B\"\n. {CHR ''a'' ^^^ n}" + assume as: "regular A" + def B\"\n. {CHR ''a'' ^^^ n}" + have b: "infinite B" + unfolding infinite_iff_countable_subset + apply(rule_tac x="\n. CHR ''a'' ^^^ n" in exI) + apply(auto simp add: inj_on_def B_def) + done + moreover + have "\x \ B. \y \ B. x \ y \ \ (x \A y)" + apply(auto simp add: B_def A_def) + apply(auto simp add: str_eq_def) + apply(drule_tac x="CHR ''b'' ^^^ aa" in spec) + (*apply(auto simp add: f_def dest: l3)*) + apply(auto) + apply(drule_tac a="CHR ''a''" and b="CHR ''b''" in length_test) + apply(simp add: length_test_def) + done + ultimately + show "False" using continuation_lemma[OF as] by blast +qed end \ No newline at end of file