diff -r c4893e84c88e -r 2455db3b06ac Myhill_2.thy --- a/Myhill_2.thy Wed Aug 03 00:52:41 2011 +0000 +++ b/Myhill_2.thy Wed Aug 03 12:36:23 2011 +0000 @@ -217,7 +217,7 @@ subsubsection {* The inductive case for @{text "Times"} *} definition - "Partitions s \ {(u, v). u @ v = s}" + "Partitions x \ {(x\<^isub>p, x\<^isub>s). x\<^isub>p @ x\<^isub>s = x}" lemma conc_partitions_elim: assumes "x \ A \ B" @@ -241,25 +241,13 @@ apply(metis append_Nil2) done - -abbreviation - tag_Times_1 :: "'a lang \ 'a lang \ 'a list \ 'a lang" +definition + tag_Times :: "'a lang \ 'a lang \ 'a list \ 'a lang \ 'a lang set" where - "tag_Times_1 A B \ \x. \A `` {x}" - -abbreviation - tag_Times_2 :: "'a lang \ 'a lang \ 'a list \ ('a lang \ 'a lang) set" -where - "tag_Times_2 A B \ \x. {(\A `` {u}, \B `` {v}) | u v. (u, v) \ Partitions x}" - -definition - tag_Times :: "'a lang \ 'a lang \ 'a list \ 'a lang \ ('a lang \ 'a lang) set" -where - "tag_Times A B \ \x. (tag_Times_1 A B x, tag_Times_2 A B x)" + "tag_Times A B \ \x. (\A `` {x}, {(\B `` {x\<^isub>s}) | x\<^isub>p x\<^isub>s. x\<^isub>p \ A \ (x\<^isub>p, x\<^isub>s) \ Partitions x})" lemma tag_Times_injI: - assumes a: "tag_Times_1 A B x = tag_Times_1 A B y" - and b: "tag_Times_2 A B x = tag_Times_2 A B y" + assumes a: "tag_Times A B x = tag_Times A B y" and c: "x @ z \ A \ B" shows "y @ z \ A \ B" proof - @@ -273,28 +261,30 @@ by (auto simp add: append_eq_append_conv2) moreover { assume eq: "x = u @ us" "us @ z = v" - have "(\A `` {u}, \B `` {us}) \ tag_Times_2 A B x" - unfolding Partitions_def using eq by (auto simp add: str_eq_def) - then have "(\A `` {u}, \B `` {us}) \ tag_Times_2 A B y" - using b by simp + have "(\B `` {us}) \ snd (tag_Times A B x)" + unfolding Partitions_def tag_Times_def using h2 eq + by (auto simp add: str_eq_def) + then have "(\B `` {us}) \ snd (tag_Times A B y)" + using a by simp then obtain u' us' where - q1: "\A `` {u} = \A `` {u'}" and + q1: "u' \ A" and q2: "\B `` {us} = \B `` {us'}" and - q3: "(u', us') \ Partitions y" by auto - from q1 h2 have "u' \ A" - using equiv_class_member by auto - moreover from q2 h3 eq + q3: "(u', us') \ Partitions y" + unfolding tag_Times_def by auto + from q2 h3 eq have "us' @ z \ B" unfolding Image_def str_eq_def by auto - ultimately have "y @ z \ A \ B" using q3 + then have "y @ z \ A \ B" using q1 q3 unfolding Partitions_def by auto } moreover { assume eq: "x @ us = u" "z = us @ v" - have "(\A `` {x}) = tag_Times_1 A B x" by simp - then have "(\A `` {x}) = tag_Times_1 A B y" + have "(\A `` {x}) = fst (tag_Times A B x)" + by (simp add: tag_Times_def) + then have "(\A `` {x}) = fst (tag_Times A B y)" using a by simp - then have "\A `` {x} = \A `` {y}" by simp + then have "\A `` {x} = \A `` {y}" + by (simp add: tag_Times_def) moreover have "x @ us \ A" using h2 eq by simp ultimately @@ -319,7 +309,7 @@ by (rule refined_intro) (auto simp add: tag_eq_def) next - have *: "finite ((UNIV // \A) \ (Pow (UNIV // \A \ UNIV // \B)))" + have *: "finite ((UNIV // \A) \ (Pow (UNIV // \B)))" using fin1 fin2 by auto show "finite (range (tag_Times A B))" unfolding tag_Times_def