diff --git a/src/HOL/Equiv_Relations.thy b/src/HOL/Equiv_Relations.thy --- a/src/HOL/Equiv_Relations.thy +++ b/src/HOL/Equiv_Relations.thy @@ -1,639 +1,642 @@ (* Title: HOL/Equiv_Relations.thy Author: Lawrence C Paulson, 1996 Cambridge University Computer Laboratory *) section \Equivalence Relations in Higher-Order Set Theory\ theory Equiv_Relations imports Groups_Big begin subsection \Equivalence relations -- set version\ definition equiv :: "'a set \ ('a \ 'a) set \ bool" where "equiv A r \ refl_on A r \ sym r \ trans r" lemma equivI: "refl_on A r \ sym r \ trans r \ equiv A r" by (simp add: equiv_def) lemma equivE: assumes "equiv A r" obtains "refl_on A r" and "sym r" and "trans r" using assms by (simp add: equiv_def) text \ Suppes, Theorem 70: \r\ is an equiv relation iff \r\ O r = r\. First half: \equiv A r \ r\ O r = r\. \ lemma sym_trans_comp_subset: "sym r \ trans r \ r\ O r \ r" unfolding trans_def sym_def converse_unfold by blast lemma refl_on_comp_subset: "refl_on A r \ r \ r\ O r" unfolding refl_on_def by blast lemma equiv_comp_eq: "equiv A r \ r\ O r = r" unfolding equiv_def by (iprover intro: sym_trans_comp_subset refl_on_comp_subset equalityI) text \Second half.\ lemma comp_equivI: assumes "r\ O r = r" "Domain r = A" shows "equiv A r" proof - have *: "\x y. (x, y) \ r \ (y, x) \ r" using assms by blast show ?thesis unfolding equiv_def refl_on_def sym_def trans_def using assms by (auto intro: *) qed subsection \Equivalence classes\ lemma equiv_class_subset: "equiv A r \ (a, b) \ r \ r``{a} \ r``{b}" \ \lemma for the next result\ unfolding equiv_def trans_def sym_def by blast theorem equiv_class_eq: "equiv A r \ (a, b) \ r \ r``{a} = r``{b}" by (intro equalityI equiv_class_subset; force simp add: equiv_def sym_def) lemma equiv_class_self: "equiv A r \ a \ A \ a \ r``{a}" unfolding equiv_def refl_on_def by blast lemma subset_equiv_class: "equiv A r \ r``{b} \ r``{a} \ b \ A \ (a, b) \ r" \ \lemma for the next result\ unfolding equiv_def refl_on_def by blast lemma eq_equiv_class: "r``{a} = r``{b} \ equiv A r \ b \ A \ (a, b) \ r" by (iprover intro: equalityD2 subset_equiv_class) lemma equiv_class_nondisjoint: "equiv A r \ x \ (r``{a} \ r``{b}) \ (a, b) \ r" unfolding equiv_def trans_def sym_def by blast lemma equiv_type: "equiv A r \ r \ A \ A" unfolding equiv_def refl_on_def by blast lemma equiv_class_eq_iff: "equiv A r \ (x, y) \ r \ r``{x} = r``{y} \ x \ A \ y \ A" by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) lemma eq_equiv_class_iff: "equiv A r \ x \ A \ y \ A \ r``{x} = r``{y} \ (x, y) \ r" by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) +lemma disjnt_equiv_class: "equiv A r \ disjnt (r``{a}) (r``{b}) \ (a, b) \ r" + by (auto dest: equiv_class_self simp: equiv_class_eq_iff disjnt_def) + subsection \Quotients\ definition quotient :: "'a set \ ('a \ 'a) set \ 'a set set" (infixl "'/'/" 90) where "A//r = (\x \ A. {r``{x}})" \ \set of equiv classes\ lemma quotientI: "x \ A \ r``{x} \ A//r" unfolding quotient_def by blast lemma quotientE: "X \ A//r \ (\x. X = r``{x} \ x \ A \ P) \ P" unfolding quotient_def by blast lemma Union_quotient: "equiv A r \ \(A//r) = A" unfolding equiv_def refl_on_def quotient_def by blast lemma quotient_disj: "equiv A r \ X \ A//r \ Y \ A//r \ X = Y \ X \ Y = {}" unfolding quotient_def equiv_def trans_def sym_def by blast lemma quotient_eqI: assumes "equiv A r" "X \ A//r" "Y \ A//r" and xy: "x \ X" "y \ Y" "(x, y) \ r" shows "X = Y" proof - obtain a b where "a \ A" and a: "X = r `` {a}" and "b \ A" and b: "Y = r `` {b}" using assms by (auto elim!: quotientE) then have "(a,b) \ r" using xy \equiv A r\ unfolding equiv_def sym_def trans_def by blast then show ?thesis unfolding a b by (rule equiv_class_eq [OF \equiv A r\]) qed lemma quotient_eq_iff: assumes "equiv A r" "X \ A//r" "Y \ A//r" and xy: "x \ X" "y \ Y" shows "X = Y \ (x, y) \ r" proof assume L: "X = Y" with assms show "(x, y) \ r" unfolding equiv_def sym_def trans_def by (blast elim!: quotientE) next assume \
: "(x, y) \ r" show "X = Y" by (rule quotient_eqI) (use \
assms in \blast+\) qed lemma eq_equiv_class_iff2: "equiv A r \ x \ A \ y \ A \ {x}//r = {y}//r \ (x, y) \ r" by (simp add: quotient_def eq_equiv_class_iff) lemma quotient_empty [simp]: "{}//r = {}" by (simp add: quotient_def) lemma quotient_is_empty [iff]: "A//r = {} \ A = {}" by (simp add: quotient_def) lemma quotient_is_empty2 [iff]: "{} = A//r \ A = {}" by (simp add: quotient_def) lemma singleton_quotient: "{x}//r = {r `` {x}}" by (simp add: quotient_def) lemma quotient_diff1: "inj_on (\a. {a}//r) A \ a \ A \ (A - {a})//r = A//r - {a}//r" unfolding quotient_def inj_on_def by blast subsection \Refinement of one equivalence relation WRT another\ lemma refines_equiv_class_eq: "R \ S \ equiv A R \ equiv A S \ R``(S``{a}) = S``{a}" by (auto simp: equiv_class_eq_iff) lemma refines_equiv_class_eq2: "R \ S \ equiv A R \ equiv A S \ S``(R``{a}) = S``{a}" by (auto simp: equiv_class_eq_iff) lemma refines_equiv_image_eq: "R \ S \ equiv A R \ equiv A S \ (\X. S``X) ` (A//R) = A//S" by (auto simp: quotient_def image_UN refines_equiv_class_eq2) lemma finite_refines_finite: "finite (A//R) \ R \ S \ equiv A R \ equiv A S \ finite (A//S)" by (erule finite_surj [where f = "\X. S``X"]) (simp add: refines_equiv_image_eq) lemma finite_refines_card_le: "finite (A//R) \ R \ S \ equiv A R \ equiv A S \ card (A//S) \ card (A//R)" by (subst refines_equiv_image_eq [of R S A, symmetric]) (auto simp: card_image_le [where f = "\X. S``X"]) subsection \Defining unary operations upon equivalence classes\ text \A congruence-preserving function.\ definition congruent :: "('a \ 'a) set \ ('a \ 'b) \ bool" where "congruent r f \ (\(y, z) \ r. f y = f z)" lemma congruentI: "(\y z. (y, z) \ r \ f y = f z) \ congruent r f" by (auto simp add: congruent_def) lemma congruentD: "congruent r f \ (y, z) \ r \ f y = f z" by (auto simp add: congruent_def) abbreviation RESPECTS :: "('a \ 'b) \ ('a \ 'a) set \ bool" (infixr "respects" 80) where "f respects r \ congruent r f" lemma UN_constant_eq: "a \ A \ \y \ A. f y = c \ (\y \ A. f y) = c" \ \lemma required to prove \UN_equiv_class\\ by auto lemma UN_equiv_class: assumes "equiv A r" "f respects r" "a \ A" shows "(\x \ r``{a}. f x) = f a" \ \Conversion rule\ proof - have \
: "\x\r `` {a}. f x = f a" using assms unfolding equiv_def congruent_def sym_def by blast show ?thesis by (iprover intro: assms UN_constant_eq [OF equiv_class_self \
]) qed lemma UN_equiv_class_type: assumes r: "equiv A r" "f respects r" and X: "X \ A//r" and AB: "\x. x \ A \ f x \ B" shows "(\x \ X. f x) \ B" using assms unfolding quotient_def by (auto simp: UN_equiv_class [OF r]) text \ Sufficient conditions for injectiveness. Could weaken premises! major premise could be an inclusion; \bcong\ could be \\y. y \ A \ f y \ B\. \ lemma UN_equiv_class_inject: assumes "equiv A r" "f respects r" and eq: "(\x \ X. f x) = (\y \ Y. f y)" and X: "X \ A//r" and Y: "Y \ A//r" and fr: "\x y. x \ A \ y \ A \ f x = f y \ (x, y) \ r" shows "X = Y" proof - obtain a b where "a \ A" and a: "X = r `` {a}" and "b \ A" and b: "Y = r `` {b}" using assms by (auto elim!: quotientE) then have "\ (f ` r `` {a}) = f a" "\ (f ` r `` {b}) = f b" by (iprover intro: UN_equiv_class [OF \equiv A r\] assms)+ then have "f a = f b" using eq unfolding a b by (iprover intro: trans sym) then have "(a,b) \ r" using fr \a \ A\ \b \ A\ by blast then show ?thesis unfolding a b by (rule equiv_class_eq [OF \equiv A r\]) qed subsection \Defining binary operations upon equivalence classes\ text \A congruence-preserving function of two arguments.\ definition congruent2 :: "('a \ 'a) set \ ('b \ 'b) set \ ('a \ 'b \ 'c) \ bool" where "congruent2 r1 r2 f \ (\(y1, z1) \ r1. \(y2, z2) \ r2. f y1 y2 = f z1 z2)" lemma congruent2I': assumes "\y1 z1 y2 z2. (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" shows "congruent2 r1 r2 f" using assms by (auto simp add: congruent2_def) lemma congruent2D: "congruent2 r1 r2 f \ (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" by (auto simp add: congruent2_def) text \Abbreviation for the common case where the relations are identical.\ abbreviation RESPECTS2:: "('a \ 'a \ 'b) \ ('a \ 'a) set \ bool" (infixr "respects2" 80) where "f respects2 r \ congruent2 r r f" lemma congruent2_implies_congruent: "equiv A r1 \ congruent2 r1 r2 f \ a \ A \ congruent r2 (f a)" unfolding congruent_def congruent2_def equiv_def refl_on_def by blast lemma congruent2_implies_congruent_UN: assumes "equiv A1 r1" "equiv A2 r2" "congruent2 r1 r2 f" "a \ A2" shows "congruent r1 (\x1. \x2 \ r2``{a}. f x1 x2)" unfolding congruent_def proof clarify fix c d assume cd: "(c,d) \ r1" then have "c \ A1" "d \ A1" using \equiv A1 r1\ by (auto elim!: equiv_type [THEN subsetD, THEN SigmaE2]) with assms show "\ (f c ` r2 `` {a}) = \ (f d ` r2 `` {a})" proof (simp add: UN_equiv_class congruent2_implies_congruent) show "f c a = f d a" using assms cd unfolding congruent2_def equiv_def refl_on_def by blast qed qed lemma UN_equiv_class2: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ a1 \ A1 \ a2 \ A2 \ (\x1 \ r1``{a1}. \x2 \ r2``{a2}. f x1 x2) = f a1 a2" by (simp add: UN_equiv_class congruent2_implies_congruent congruent2_implies_congruent_UN) lemma UN_equiv_class_type2: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ X1 \ A1//r1 \ X2 \ A2//r2 \ (\x1 x2. x1 \ A1 \ x2 \ A2 \ f x1 x2 \ B) \ (\x1 \ X1. \x2 \ X2. f x1 x2) \ B" unfolding quotient_def by (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN congruent2_implies_congruent quotientI) lemma UN_UN_split_split_eq: "(\(x1, x2) \ X. \(y1, y2) \ Y. A x1 x2 y1 y2) = (\x \ X. \y \ Y. (\(x1, x2). (\(y1, y2). A x1 x2 y1 y2) y) x)" \ \Allows a natural expression of binary operators,\ \ \without explicit calls to \split\\ by auto lemma congruent2I: "equiv A1 r1 \ equiv A2 r2 \ (\y z w. w \ A2 \ (y,z) \ r1 \ f y w = f z w) \ (\y z w. w \ A1 \ (y,z) \ r2 \ f w y = f w z) \ congruent2 r1 r2 f" \ \Suggested by John Harrison -- the two subproofs may be\ \ \\<^emph>\much\ simpler than the direct proof.\ unfolding congruent2_def equiv_def refl_on_def by (blast intro: trans) lemma congruent2_commuteI: assumes equivA: "equiv A r" and commute: "\y z. y \ A \ z \ A \ f y z = f z y" and congt: "\y z w. w \ A \ (y,z) \ r \ f w y = f w z" shows "f respects2 r" proof (rule congruent2I [OF equivA equivA]) note eqv = equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2] show "\y z w. \w \ A; (y, z) \ r\ \ f y w = f z w" by (iprover intro: commute [THEN trans] sym congt elim: eqv) show "\y z w. \w \ A; (y, z) \ r\ \ f w y = f w z" by (iprover intro: congt elim: eqv) qed subsection \Quotients and finiteness\ text \Suggested by Florian Kammüller\ lemma finite_quotient: assumes "finite A" "r \ A \ A" shows "finite (A//r)" \ \recall @{thm equiv_type}\ proof - have "A//r \ Pow A" using assms unfolding quotient_def by blast moreover have "finite (Pow A)" using assms by simp ultimately show ?thesis by (iprover intro: finite_subset) qed lemma finite_equiv_class: "finite A \ r \ A \ A \ X \ A//r \ finite X" unfolding quotient_def by (erule rev_finite_subset) blast lemma equiv_imp_dvd_card: assumes "finite A" "equiv A r" "\X. X \ A//r \ k dvd card X" shows "k dvd card A" proof (rule Union_quotient [THEN subst]) show "k dvd card (\ (A // r))" apply (rule dvd_partition) using assms by (auto simp: Union_quotient dest: quotient_disj) qed (use assms in blast) lemma card_quotient_disjoint: assumes "finite A" "inj_on (\x. {x} // r) A" shows "card (A//r) = card A" proof - have "\i\A. \j\A. i \ j \ r `` {j} \ r `` {i}" using assms by (fastforce simp add: quotient_def inj_on_def) with assms show ?thesis by (simp add: quotient_def card_UN_disjoint) qed text \By Jakub Kądziołka:\ lemma sum_fun_comp: assumes "finite S" "finite R" "g ` S \ R" shows "(\x \ S. f (g x)) = (\y \ R. of_nat (card {x \ S. g x = y}) * f y)" proof - let ?r = "relation_of (\p q. g p = g q) S" have eqv: "equiv S ?r" unfolding relation_of_def by (auto intro: comp_equivI) have finite: "C \ S//?r \ finite C" for C by (fact finite_equiv_class[OF `finite S` equiv_type[OF `equiv S ?r`]]) have disjoint: "A \ S//?r \ B \ S//?r \ A \ B \ A \ B = {}" for A B using eqv quotient_disj by blast let ?cls = "\y. {x \ S. y = g x}" have quot_as_img: "S//?r = ?cls ` g ` S" by (auto simp add: relation_of_def quotient_def) have cls_inj: "inj_on ?cls (g ` S)" by (auto intro: inj_onI) have rest_0: "(\y \ R - g ` S. of_nat (card (?cls y)) * f y) = 0" proof - have "of_nat (card (?cls y)) * f y = 0" if asm: "y \ R - g ` S" for y proof - from asm have *: "?cls y = {}" by auto show ?thesis unfolding * by simp qed thus ?thesis by simp qed have "(\x \ S. f (g x)) = (\C \ S//?r. \x \ C. f (g x))" using eqv finite disjoint by (simp flip: sum.Union_disjoint[simplified] add: Union_quotient) also have "... = (\y \ g ` S. \x \ ?cls y. f (g x))" unfolding quot_as_img by (simp add: sum.reindex[OF cls_inj]) also have "... = (\y \ g ` S. \x \ ?cls y. f y)" by auto also have "... = (\y \ g ` S. of_nat (card (?cls y)) * f y)" by (simp flip: sum_constant) also have "... = (\y \ R. of_nat (card (?cls y)) * f y)" using rest_0 by (simp add: sum.subset_diff[OF \g ` S \ R\ \finite R\]) finally show ?thesis by (simp add: eq_commute) qed subsection \Projection\ definition proj :: "('b \ 'a) set \ 'b \ 'a set" where "proj r x = r `` {x}" lemma proj_preserves: "x \ A \ proj r x \ A//r" unfolding proj_def by (rule quotientI) lemma proj_in_iff: assumes "equiv A r" shows "proj r x \ A//r \ x \ A" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: proj_preserves) next assume ?lhs then show ?rhs unfolding proj_def quotient_def proof clarsimp fix y assume y: "y \ A" and "r `` {x} = r `` {y}" moreover have "y \ r `` {y}" using assms y unfolding equiv_def refl_on_def by blast ultimately have "(x, y) \ r" by blast then show "x \ A" using assms unfolding equiv_def refl_on_def by blast qed qed lemma proj_iff: "equiv A r \ {x, y} \ A \ proj r x = proj r y \ (x, y) \ r" by (simp add: proj_def eq_equiv_class_iff) (* lemma in_proj: "\equiv A r; x \ A\ \ x \ proj r x" unfolding proj_def equiv_def refl_on_def by blast *) lemma proj_image: "proj r ` A = A//r" unfolding proj_def[abs_def] quotient_def by blast lemma in_quotient_imp_non_empty: "equiv A r \ X \ A//r \ X \ {}" unfolding quotient_def using equiv_class_self by fast lemma in_quotient_imp_in_rel: "equiv A r \ X \ A//r \ {x, y} \ X \ (x, y) \ r" using quotient_eq_iff[THEN iffD1] by fastforce lemma in_quotient_imp_closed: "equiv A r \ X \ A//r \ x \ X \ (x, y) \ r \ y \ X" unfolding quotient_def equiv_def trans_def by blast lemma in_quotient_imp_subset: "equiv A r \ X \ A//r \ X \ A" using in_quotient_imp_in_rel equiv_type by fastforce subsection \Equivalence relations -- predicate version\ text \Partial equivalences.\ definition part_equivp :: "('a \ 'a \ bool) \ bool" where "part_equivp R \ (\x. R x x) \ (\x y. R x y \ R x x \ R y y \ R x = R y)" \ \John-Harrison-style characterization\ lemma part_equivpI: "\x. R x x \ symp R \ transp R \ part_equivp R" by (auto simp add: part_equivp_def) (auto elim: sympE transpE) lemma part_equivpE: assumes "part_equivp R" obtains x where "R x x" and "symp R" and "transp R" proof - from assms have 1: "\x. R x x" and 2: "\x y. R x y \ R x x \ R y y \ R x = R y" unfolding part_equivp_def by blast+ from 1 obtain x where "R x x" .. moreover have "symp R" proof (rule sympI) fix x y assume "R x y" with 2 [of x y] show "R y x" by auto qed moreover have "transp R" proof (rule transpI) fix x y z assume "R x y" and "R y z" with 2 [of x y] 2 [of y z] show "R x z" by auto qed ultimately show thesis by (rule that) qed lemma part_equivp_refl_symp_transp: "part_equivp R \ (\x. R x x) \ symp R \ transp R" by (auto intro: part_equivpI elim: part_equivpE) lemma part_equivp_symp: "part_equivp R \ R x y \ R y x" by (erule part_equivpE, erule sympE) lemma part_equivp_transp: "part_equivp R \ R x y \ R y z \ R x z" by (erule part_equivpE, erule transpE) lemma part_equivp_typedef: "part_equivp R \ \d. d \ {c. \x. R x x \ c = Collect (R x)}" by (auto elim: part_equivpE) text \Total equivalences.\ definition equivp :: "('a \ 'a \ bool) \ bool" where "equivp R \ (\x y. R x y = (R x = R y))" \ \John-Harrison-style characterization\ lemma equivpI: "reflp R \ symp R \ transp R \ equivp R" by (auto elim: reflpE sympE transpE simp add: equivp_def) lemma equivpE: assumes "equivp R" obtains "reflp R" and "symp R" and "transp R" using assms by (auto intro!: that reflpI sympI transpI simp add: equivp_def) lemma equivp_implies_part_equivp: "equivp R \ part_equivp R" by (auto intro: part_equivpI elim: equivpE reflpE) lemma equivp_equiv: "equiv UNIV A \ equivp (\x y. (x, y) \ A)" by (auto intro!: equivI equivpI [to_set] elim!: equivE equivpE [to_set]) lemma equivp_reflp_symp_transp: "equivp R \ reflp R \ symp R \ transp R" by (auto intro: equivpI elim: equivpE) lemma identity_equivp: "equivp (=)" by (auto intro: equivpI reflpI sympI transpI) lemma equivp_reflp: "equivp R \ R x x" by (erule equivpE, erule reflpE) lemma equivp_symp: "equivp R \ R x y \ R y x" by (erule equivpE, erule sympE) lemma equivp_transp: "equivp R \ R x y \ R y z \ R x z" by (erule equivpE, erule transpE) lemma equivp_rtranclp: "symp r \ equivp r\<^sup>*\<^sup>*" by(intro equivpI reflpI sympI transpI)(auto dest: sympD[OF symp_rtranclp]) lemmas equivp_rtranclp_symclp [simp] = equivp_rtranclp[OF symp_symclp] lemma equivp_vimage2p: "equivp R \ equivp (vimage2p f f R)" by(auto simp add: equivp_def vimage2p_def dest: fun_cong) lemma equivp_imp_transp: "equivp R \ transp R" by(simp add: equivp_reflp_symp_transp) subsection \Equivalence closure\ definition equivclp :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" where "equivclp r = (symclp r)\<^sup>*\<^sup>*" lemma transp_equivclp [simp]: "transp (equivclp r)" by(simp add: equivclp_def) lemma reflp_equivclp [simp]: "reflp (equivclp r)" by(simp add: equivclp_def) lemma symp_equivclp [simp]: "symp (equivclp r)" by(simp add: equivclp_def) lemma equivp_evquivclp [simp]: "equivp (equivclp r)" by(simp add: equivpI) lemma tranclp_equivclp [simp]: "(equivclp r)\<^sup>+\<^sup>+ = equivclp r" by(simp add: equivclp_def) lemma rtranclp_equivclp [simp]: "(equivclp r)\<^sup>*\<^sup>* = equivclp r" by(simp add: equivclp_def) lemma symclp_equivclp [simp]: "symclp (equivclp r) = equivclp r" by(simp add: equivclp_def symp_symclp_eq) lemma equivclp_symclp [simp]: "equivclp (symclp r) = equivclp r" by(simp add: equivclp_def) lemma equivclp_conversep [simp]: "equivclp (conversep r) = equivclp r" by(simp add: equivclp_def) lemma equivclp_sym [sym]: "equivclp r x y \ equivclp r y x" by(rule sympD[OF symp_equivclp]) lemma equivclp_OO_equivclp_le_equivclp: "equivclp r OO equivclp r \ equivclp r" by(rule transp_relcompp_less_eq transp_equivclp)+ lemma rtranlcp_le_equivclp: "r\<^sup>*\<^sup>* \ equivclp r" unfolding equivclp_def by(rule rtranclp_mono)(simp add: symclp_pointfree) lemma rtranclp_conversep_le_equivclp: "r\\\<^sup>*\<^sup>* \ equivclp r" unfolding equivclp_def by(rule rtranclp_mono)(simp add: symclp_pointfree) lemma symclp_rtranclp_le_equivclp: "symclp r\<^sup>*\<^sup>* \ equivclp r" unfolding symclp_pointfree by(rule le_supI)(simp_all add: rtranclp_conversep[symmetric] rtranlcp_le_equivclp rtranclp_conversep_le_equivclp) lemma r_OO_conversep_into_equivclp: "r\<^sup>*\<^sup>* OO r\\\<^sup>*\<^sup>* \ equivclp r" by(blast intro: order_trans[OF _ equivclp_OO_equivclp_le_equivclp] relcompp_mono rtranlcp_le_equivclp rtranclp_conversep_le_equivclp del: predicate2I) lemma equivclp_induct [consumes 1, case_names base step, induct pred: equivclp]: assumes a: "equivclp r a b" and cases: "P a" "\y z. equivclp r a y \ r y z \ r z y \ P y \ P z" shows "P b" using a unfolding equivclp_def by(induction rule: rtranclp_induct; fold equivclp_def; blast intro: cases elim: symclpE) lemma converse_equivclp_induct [consumes 1, case_names base step]: assumes major: "equivclp r a b" and cases: "P b" "\y z. r y z \ r z y \ equivclp r z b \ P z \ P y" shows "P a" using major unfolding equivclp_def by(induction rule: converse_rtranclp_induct; fold equivclp_def; blast intro: cases elim: symclpE) lemma equivclp_refl [simp]: "equivclp r x x" by(rule reflpD[OF reflp_equivclp]) lemma r_into_equivclp [intro]: "r x y \ equivclp r x y" unfolding equivclp_def by(blast intro: symclpI) lemma converse_r_into_equivclp [intro]: "r y x \ equivclp r x y" unfolding equivclp_def by(blast intro: symclpI) lemma rtranclp_into_equivclp: "r\<^sup>*\<^sup>* x y \ equivclp r x y" using rtranlcp_le_equivclp[of r] by blast lemma converse_rtranclp_into_equivclp: "r\<^sup>*\<^sup>* y x \ equivclp r x y" by(blast intro: equivclp_sym rtranclp_into_equivclp) lemma equivclp_into_equivclp: "\ equivclp r a b; r b c \ r c b \ \ equivclp r a c" unfolding equivclp_def by(erule rtranclp.rtrancl_into_rtrancl)(auto intro: symclpI) lemma equivclp_trans [trans]: "\ equivclp r a b; equivclp r b c \ \ equivclp r a c" using equivclp_OO_equivclp_le_equivclp[of r] by blast hide_const (open) proj end diff --git a/src/HOL/Library/Disjoint_Sets.thy b/src/HOL/Library/Disjoint_Sets.thy --- a/src/HOL/Library/Disjoint_Sets.thy +++ b/src/HOL/Library/Disjoint_Sets.thy @@ -1,499 +1,598 @@ (* Author: Johannes Hölzl, TU München *) section \Partitions and Disjoint Sets\ theory Disjoint_Sets - imports Main + imports FuncSet begin -lemma range_subsetD: "range f \ B \ f i \ B" - by blast - -lemma Int_Diff_disjoint: "A \ B \ (A - B) = {}" - by blast - -lemma Int_Diff_Un: "A \ B \ (A - B) = A" - by blast - -lemma mono_Un: "mono A \ (\i\n. A i) = A n" +lemma mono_imp_UN_eq_last: "mono A \ (\i\n. A i) = A n" unfolding mono_def by auto -lemma disjnt_equiv_class: "equiv A r \ disjnt (r``{a}) (r``{b}) \ (a, b) \ r" - by (auto dest: equiv_class_self simp: equiv_class_eq_iff disjnt_def) - subsection \Set of Disjoint Sets\ abbreviation disjoint :: "'a set set \ bool" where "disjoint \ pairwise disjnt" lemma disjoint_def: "disjoint A \ (\a\A. \b\A. a \ b \ a \ b = {})" unfolding pairwise_def disjnt_def by auto lemma disjointI: "(\a b. a \ A \ b \ A \ a \ b \ a \ b = {}) \ disjoint A" unfolding disjoint_def by auto lemma disjointD: "disjoint A \ a \ A \ b \ A \ a \ b \ a \ b = {}" unfolding disjoint_def by auto lemma disjoint_image: "inj_on f (\A) \ disjoint A \ disjoint ((`) f ` A)" unfolding inj_on_def disjoint_def by blast lemma assumes "disjoint (A \ B)" shows disjoint_unionD1: "disjoint A" and disjoint_unionD2: "disjoint B" using assms by (simp_all add: disjoint_def) lemma disjoint_INT: assumes *: "\i. i \ I \ disjoint (F i)" shows "disjoint {\i\I. X i | X. \i\I. X i \ F i}" proof (safe intro!: disjointI del: equalityI) fix A B :: "'a \ 'b set" assume "(\i\I. A i) \ (\i\I. B i)" then obtain i where "A i \ B i" "i \ I" by auto moreover assume "\i\I. A i \ F i" "\i\I. B i \ F i" ultimately show "(\i\I. A i) \ (\i\I. B i) = {}" using *[OF \i\I\, THEN disjointD, of "A i" "B i"] by (auto simp flip: INT_Int_distrib) qed lemma diff_Union_pairwise_disjoint: assumes "pairwise disjnt \" "\ \ \" shows "\\ - \\ = \(\ - \)" proof - have "False" if x: "x \ A" "x \ B" and AB: "A \ \" "A \ \" "B \ \" for x A B proof - have "A \ B = {}" using assms disjointD AB by blast with x show ?thesis by blast qed then show ?thesis by auto qed lemma Int_Union_pairwise_disjoint: assumes "pairwise disjnt (\ \ \)" shows "\\ \ \\ = \(\ \ \)" proof - have "False" if x: "x \ A" "x \ B" and AB: "A \ \" "A \ \" "B \ \" for x A B proof - have "A \ B = {}" using assms disjointD AB by blast with x show ?thesis by blast qed then show ?thesis by auto qed lemma psubset_Union_pairwise_disjoint: assumes \: "pairwise disjnt \" and "\ \ \ - {{}}" shows "\\ \ \\" unfolding psubset_eq proof show "\\ \ \\" using assms by blast have "\ \ \" "\(\ - \ \ (\ - {{}})) \ {}" using assms by blast+ then show "\\ \ \\" using diff_Union_pairwise_disjoint [OF \] by blast qed subsubsection "Family of Disjoint Sets" definition disjoint_family_on :: "('i \ 'a set) \ 'i set \ bool" where "disjoint_family_on A S \ (\m\S. \n\S. m \ n \ A m \ A n = {})" abbreviation "disjoint_family A \ disjoint_family_on A UNIV" lemma disjoint_family_elem_disjnt: assumes "infinite A" "finite C" and df: "disjoint_family_on B A" obtains x where "x \ A" "disjnt C (B x)" proof - have "False" if *: "\x \ A. \y. y \ C \ y \ B x" proof - obtain g where g: "\x \ A. g x \ C \ g x \ B x" using * by metis with df have "inj_on g A" by (fastforce simp add: inj_on_def disjoint_family_on_def) then have "infinite (g ` A)" using \infinite A\ finite_image_iff by blast then show False by (meson \finite C\ finite_subset g image_subset_iff) qed then show ?thesis by (force simp: disjnt_iff intro: that) qed lemma disjoint_family_onD: "disjoint_family_on A I \ i \ I \ j \ I \ i \ j \ A i \ A j = {}" by (auto simp: disjoint_family_on_def) lemma disjoint_family_subset: "disjoint_family A \ (\x. B x \ A x) \ disjoint_family B" by (force simp add: disjoint_family_on_def) lemma disjoint_family_on_bisimulation: assumes "disjoint_family_on f S" and "\n m. n \ S \ m \ S \ n \ m \ f n \ f m = {} \ g n \ g m = {}" shows "disjoint_family_on g S" using assms unfolding disjoint_family_on_def by auto lemma disjoint_family_on_mono: "A \ B \ disjoint_family_on f B \ disjoint_family_on f A" unfolding disjoint_family_on_def by auto lemma disjoint_family_Suc: "(\n. A n \ A (Suc n)) \ disjoint_family (\i. A (Suc i) - A i)" using lift_Suc_mono_le[of A] by (auto simp add: disjoint_family_on_def) (metis insert_absorb insert_subset le_SucE le_antisym not_le_imp_less less_imp_le) lemma disjoint_family_on_disjoint_image: "disjoint_family_on A I \ disjoint (A ` I)" unfolding disjoint_family_on_def disjoint_def by force lemma disjoint_family_on_vimageI: "disjoint_family_on F I \ disjoint_family_on (\i. f -` F i) I" by (auto simp: disjoint_family_on_def) lemma disjoint_image_disjoint_family_on: assumes d: "disjoint (A ` I)" and i: "inj_on A I" shows "disjoint_family_on A I" unfolding disjoint_family_on_def proof (intro ballI impI) fix n m assume nm: "m \ I" "n \ I" and "n \ m" with i[THEN inj_onD, of n m] show "A n \ A m = {}" by (intro disjointD[OF d]) auto qed lemma disjoint_family_on_iff_disjoint_image: assumes "\i. i \ I \ A i \ {}" shows "disjoint_family_on A I \ disjoint (A ` I) \ inj_on A I" proof assume "disjoint_family_on A I" then show "disjoint (A ` I) \ inj_on A I" by (metis (mono_tags, lifting) assms disjoint_family_onD disjoint_family_on_disjoint_image inf.idem inj_onI) qed (use disjoint_image_disjoint_family_on in metis) lemma card_UN_disjoint': assumes "disjoint_family_on A I" "\i. i \ I \ finite (A i)" "finite I" shows "card (\i\I. A i) = (\i\I. card (A i))" using assms by (simp add: card_UN_disjoint disjoint_family_on_def) lemma disjoint_UN: assumes F: "\i. i \ I \ disjoint (F i)" and *: "disjoint_family_on (\i. \(F i)) I" shows "disjoint (\i\I. F i)" proof (safe intro!: disjointI del: equalityI) fix A B i j assume "A \ B" "A \ F i" "i \ I" "B \ F j" "j \ I" show "A \ B = {}" proof cases assume "i = j" with F[of i] \i \ I\ \A \ F i\ \B \ F j\ \A \ B\ show "A \ B = {}" by (auto dest: disjointD) next assume "i \ j" with * \i\I\ \j\I\ have "(\(F i)) \ (\(F j)) = {}" by (rule disjoint_family_onD) with \A\F i\ \i\I\ \B\F j\ \j\I\ show "A \ B = {}" by auto qed qed lemma distinct_list_bind: assumes "distinct xs" "\x. x \ set xs \ distinct (f x)" "disjoint_family_on (set \ f) (set xs)" shows "distinct (List.bind xs f)" using assms by (induction xs) (auto simp: disjoint_family_on_def distinct_map inj_on_def set_list_bind) lemma bij_betw_UNION_disjoint: assumes disj: "disjoint_family_on A' I" assumes bij: "\i. i \ I \ bij_betw f (A i) (A' i)" shows "bij_betw f (\i\I. A i) (\i\I. A' i)" unfolding bij_betw_def proof from bij show eq: "f ` \(A ` I) = \(A' ` I)" by (auto simp: bij_betw_def image_UN) show "inj_on f (\(A ` I))" proof (rule inj_onI, clarify) fix i j x y assume A: "i \ I" "j \ I" "x \ A i" "y \ A j" and B: "f x = f y" from A bij[of i] bij[of j] have "f x \ A' i" "f y \ A' j" by (auto simp: bij_betw_def) with B have "A' i \ A' j \ {}" by auto with disj A have "i = j" unfolding disjoint_family_on_def by blast with A B bij[of i] show "x = y" by (auto simp: bij_betw_def dest: inj_onD) qed qed lemma disjoint_union: "disjoint C \ disjoint B \ \C \ \B = {} \ disjoint (C \ B)" using disjoint_UN[of "{C, B}" "\x. x"] by (auto simp add: disjoint_family_on_def) text \ Sum/product of the union of a finite disjoint family \ context comm_monoid_set begin lemma UNION_disjoint_family: assumes "finite I" and "\i\I. finite (A i)" and "disjoint_family_on A I" shows "F g (\(A ` I)) = F (\x. F g (A x)) I" using assms unfolding disjoint_family_on_def by (rule UNION_disjoint) lemma Union_disjoint_sets: assumes "\A\C. finite A" and "disjoint C" shows "F g (\C) = (F \ F) g C" using assms unfolding disjoint_def by (rule Union_disjoint) end text \ The union of an infinite disjoint family of non-empty sets is infinite. \ lemma infinite_disjoint_family_imp_infinite_UNION: assumes "\finite A" "\x. x \ A \ f x \ {}" "disjoint_family_on f A" shows "\finite (\(f ` A))" proof - define g where "g x = (SOME y. y \ f x)" for x have g: "g x \ f x" if "x \ A" for x unfolding g_def by (rule someI_ex, insert assms(2) that) blast have inj_on_g: "inj_on g A" proof (rule inj_onI, rule ccontr) fix x y assume A: "x \ A" "y \ A" "g x = g y" "x \ y" with g[of x] g[of y] have "g x \ f x" "g x \ f y" by auto with A \x \ y\ assms show False by (auto simp: disjoint_family_on_def inj_on_def) qed from g have "g ` A \ \(f ` A)" by blast moreover from inj_on_g \\finite A\ have "\finite (g ` A)" using finite_imageD by blast ultimately show ?thesis using finite_subset by blast qed subsection \Construct Disjoint Sequences\ definition disjointed :: "(nat \ 'a set) \ nat \ 'a set" where "disjointed A n = A n - (\i\{0..i\{0..i\{0..i. disjointed A i) = (\i. A i)" by (rule UN_finite2_eq [where k=0]) (simp add: finite_UN_disjointed_eq) lemma less_disjoint_disjointed: "m < n \ disjointed A m \ disjointed A n = {}" by (auto simp add: disjointed_def) lemma disjoint_family_disjointed: "disjoint_family (disjointed A)" by (simp add: disjoint_family_on_def) (metis neq_iff Int_commute less_disjoint_disjointed) lemma disjointed_subset: "disjointed A n \ A n" by (auto simp add: disjointed_def) lemma disjointed_0[simp]: "disjointed A 0 = A 0" by (simp add: disjointed_def) lemma disjointed_mono: "mono A \ disjointed A (Suc n) = A (Suc n) - A n" - using mono_Un[of A] by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost) + using mono_imp_UN_eq_last[of A] by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost) subsection \Partitions\ text \ Partitions \<^term>\P\ of a set \<^term>\A\. We explicitly disallow empty sets. \ definition partition_on :: "'a set \ 'a set set \ bool" where "partition_on A P \ \P = A \ disjoint P \ {} \ P" lemma partition_onI: "\P = A \ (\p q. p \ P \ q \ P \ p \ q \ disjnt p q) \ {} \ P \ partition_on A P" by (auto simp: partition_on_def pairwise_def) lemma partition_onD1: "partition_on A P \ A = \P" by (auto simp: partition_on_def) lemma partition_onD2: "partition_on A P \ disjoint P" by (auto simp: partition_on_def) lemma partition_onD3: "partition_on A P \ {} \ P" by (auto simp: partition_on_def) subsection \Constructions of partitions\ lemma partition_on_empty: "partition_on {} P \ P = {}" unfolding partition_on_def by fastforce lemma partition_on_space: "A \ {} \ partition_on A {A}" by (auto simp: partition_on_def disjoint_def) lemma partition_on_singletons: "partition_on A ((\x. {x}) ` A)" by (auto simp: partition_on_def disjoint_def) lemma partition_on_transform: assumes P: "partition_on A P" assumes F_UN: "\(F ` P) = F (\P)" and F_disjnt: "\p q. p \ P \ q \ P \ disjnt p q \ disjnt (F p) (F q)" shows "partition_on (F A) (F ` P - {{}})" proof - have "\(F ` P - {{}}) = F A" unfolding P[THEN partition_onD1] F_UN[symmetric] by auto with P show ?thesis by (auto simp add: partition_on_def pairwise_def intro!: F_disjnt) qed lemma partition_on_restrict: "partition_on A P \ partition_on (B \ A) ((\) B ` P - {{}})" by (intro partition_on_transform) (auto simp: disjnt_def) lemma partition_on_vimage: "partition_on A P \ partition_on (f -` A) ((-`) f ` P - {{}})" by (intro partition_on_transform) (auto simp: disjnt_def) lemma partition_on_inj_image: assumes P: "partition_on A P" and f: "inj_on f A" shows "partition_on (f ` A) ((`) f ` P - {{}})" proof (rule partition_on_transform[OF P]) show "p \ P \ q \ P \ disjnt p q \ disjnt (f ` p) (f ` q)" for p q using f[THEN inj_onD] P[THEN partition_onD1] by (auto simp: disjnt_def) qed auto lemma partition_on_insert: assumes "disjnt p (\P)" shows "partition_on A (insert p P) \ partition_on (A-p) P \ p \ A \ p \ {}" using assms by (auto simp: partition_on_def disjnt_iff pairwise_insert) subsection \Finiteness of partitions\ lemma finitely_many_partition_on: assumes "finite A" shows "finite {P. partition_on A P}" proof (rule finite_subset) show "{P. partition_on A P} \ Pow (Pow A)" unfolding partition_on_def by auto show "finite (Pow (Pow A))" using assms by simp qed lemma finite_elements: "finite A \ partition_on A P \ finite P" using partition_onD1[of A P] by (simp add: finite_UnionD) lemma product_partition: assumes "partition_on A P" and "\p. p \ P \ finite p" shows "card A = (\p\P. card p)" using assms unfolding partition_on_def by (meson card_Union_disjoint) subsection \Equivalence of partitions and equivalence classes\ lemma partition_on_quotient: assumes r: "equiv A r" shows "partition_on A (A // r)" proof (rule partition_onI) from r have "refl_on A r" by (auto elim: equivE) then show "\(A // r) = A" "{} \ A // r" by (auto simp: refl_on_def quotient_def) fix p q assume "p \ A // r" "q \ A // r" "p \ q" then obtain x y where "x \ A" "y \ A" "p = r `` {x}" "q = r `` {y}" by (auto simp: quotient_def) with r equiv_class_eq_iff[OF r, of x y] \p \ q\ show "disjnt p q" by (auto simp: disjnt_equiv_class) qed lemma equiv_partition_on: assumes P: "partition_on A P" shows "equiv A {(x, y). \p \ P. x \ p \ y \ p}" proof (rule equivI) have "A = \P" "disjoint P" "{} \ P" using P by (auto simp: partition_on_def) then show "refl_on A {(x, y). \p\P. x \ p \ y \ p}" unfolding refl_on_def by auto show "trans {(x, y). \p\P. x \ p \ y \ p}" using \disjoint P\ by (auto simp: trans_def disjoint_def) qed (auto simp: sym_def) lemma partition_on_eq_quotient: assumes P: "partition_on A P" shows "A // {(x, y). \p \ P. x \ p \ y \ p} = P" unfolding quotient_def proof safe fix x assume "x \ A" then obtain p where "p \ P" "x \ p" "\q. q \ P \ x \ q \ p = q" using P by (auto simp: partition_on_def disjoint_def) then have "{y. \p\P. x \ p \ y \ p} = p" by (safe intro!: bexI[of _ p]) simp then show "{(x, y). \p\P. x \ p \ y \ p} `` {x} \ P" by (simp add: \p \ P\) next fix p assume "p \ P" then have "p \ {}" using P by (auto simp: partition_on_def) then obtain x where "x \ p" by auto then have "x \ A" "\q. q \ P \ x \ q \ p = q" using P \p \ P\ by (auto simp: partition_on_def disjoint_def) with \p\P\ \x \ p\ have "{y. \p\P. x \ p \ y \ p} = p" by (safe intro!: bexI[of _ p]) simp then show "p \ (\x\A. {{(x, y). \p\P. x \ p \ y \ p} `` {x}})" by (auto intro: \x \ A\) qed lemma partition_on_alt: "partition_on A P \ (\r. equiv A r \ P = A // r)" by (auto simp: partition_on_eq_quotient intro!: partition_on_quotient intro: equiv_partition_on) subsection \Refinement of partitions\ definition refines :: "'a set \ 'a set set \ 'a set set \ bool" where "refines A P Q \ partition_on A P \ partition_on A Q \ (\X\P. \Y\Q. X \ Y)" lemma refines_refl: "partition_on A P \ refines A P P" using refines_def by blast lemma refines_asym1: assumes "refines A P Q" "refines A Q P" shows "P \ Q" proof fix X assume "X \ P" then obtain Y X' where "Y \ Q" "X \ Y" "X' \ P" "Y \ X'" by (meson assms refines_def) then have "X' = X" using assms(2) unfolding partition_on_def refines_def by (metis \X \ P\ \X \ Y\ disjnt_self_iff_empty disjnt_subset1 pairwiseD) then show "X \ Q" using \X \ Y\ \Y \ Q\ \Y \ X'\ by force qed lemma refines_asym: "\refines A P Q; refines A Q P\ \ P=Q" by (meson antisym_conv refines_asym1) lemma refines_trans: "\refines A P Q; refines A Q R\ \ refines A P R" by (meson order.trans refines_def) lemma refines_obtains_subset: assumes "refines A P Q" "q \ Q" shows "partition_on q {p \ P. p \ q}" proof - have "p \ q \ disjnt p q" if "p \ P" for p using that assms unfolding refines_def partition_on_def disjoint_def by (metis disjnt_def disjnt_subset1) with assms have "q \ Union {p \ P. p \ q}" using assms by (clarsimp simp: refines_def disjnt_iff partition_on_def) (metis Union_iff) with assms have "q = Union {p \ P. p \ q}" by auto then show ?thesis using assms by (auto simp: refines_def disjoint_def partition_on_def) qed +subsection \The coarsest common refinement of a set of partitions\ + +definition common_refinement :: "'a set set set \ 'a set set" + where "common_refinement \

\ (\f \ (\\<^sub>E P\\

. P). {\ (f ` \

)}) - {{}}" + +text \With non-extensional function space\ +lemma common_refinement: "common_refinement \

= (\f \ (\ P\\

. P). {\ (f ` \

)}) - {{}}" + (is "?lhs = ?rhs") +proof + show "?rhs \ ?lhs" + apply (clarsimp simp add: common_refinement_def PiE_def Ball_def) + by (metis restrict_Pi_cancel image_restrict_eq restrict_extensional) +qed (auto simp add: common_refinement_def PiE_def) + +lemma common_refinement_exists: "\X \ common_refinement \

; P \ \

\ \ \R\P. X \ R" + by (auto simp add: common_refinement) + +lemma Union_common_refinement: "\ (common_refinement \

) = (\ P\\

. \P)" +proof + show "(\ P\\

. \P) \ \ (common_refinement \

)" + proof (clarsimp simp: common_refinement) + fix x + assume "\P\\

. \X\P. x \ X" + then obtain F where F: "\P. P\\

\ F P \ P \ x \ F P" + by metis + then have "x \ \ (F ` \

)" + by force + with F show "\X\(\x\\ P\\

. P. {\ (x ` \

)}) - {{}}. x \ X" + by (auto simp add: Pi_iff Bex_def) + qed +qed (auto simp: common_refinement_def) + +lemma partition_on_common_refinement: + assumes A: "\P. P \ \

\ partition_on A P" and "\

\ {}" + shows "partition_on A (common_refinement \

)" +proof (rule partition_onI) + show "\ (common_refinement \

) = A" + using assms by (simp add: partition_on_def Union_common_refinement) + fix P Q + assume "P \ common_refinement \

" and "Q \ common_refinement \

" and "P \ Q" + then obtain f g where f: "f \ (\\<^sub>E P\\

. P)" and P: "P = \ (f ` \

)" and "P \ {}" + and g: "g \ (\\<^sub>E P\\

. P)" and Q: "Q = \ (g ` \

)" and "Q \ {}" + by (auto simp add: common_refinement_def) + have "f=g" if "x \ P" "x \ Q" for x + proof (rule extensionalityI [of _ \

]) + fix R + assume "R \ \

" + with that P Q f g A [unfolded partition_on_def, OF \R \ \

\] + show "f R = g R" + by (metis INT_E Int_iff PiE_iff disjointD emptyE) + qed (use PiE_iff f g in auto) + then show "disjnt P Q" + by (metis P Q \P \ Q\ disjnt_iff) +qed (simp add: common_refinement_def) + +lemma refines_common_refinement: + assumes "\P. P \ \

\ partition_on A P" "P \ \

" + shows "refines A (common_refinement \

) P" + unfolding refines_def +proof (intro conjI strip) + fix X + assume "X \ common_refinement \

" + with assms show "\Y\P. X \ Y" + by (auto simp: common_refinement_def) +qed (use assms partition_on_common_refinement in auto) + +text \The common refinement is itself refined by any other\ +lemma common_refinement_coarsest: + assumes "\P. P \ \

\ partition_on A P" "partition_on A R" "\P. P \ \

\ refines A R P" "\

\ {}" + shows "refines A R (common_refinement \

)" + unfolding refines_def +proof (intro conjI ballI partition_on_common_refinement) + fix X + assume "X \ R" + have "\p \ P. X \ p" if "P \ \

" for P + by (meson \X \ R\ assms(3) refines_def that) + then obtain F where f: "\P. P \ \

\ F P \ P \ X \ F P" + by metis + with \partition_on A R\ \X \ R\ \\

\ {}\ + have "\ (F ` \

) \ common_refinement \

" + apply (simp add: partition_on_def common_refinement Pi_iff Bex_def) + by (metis (no_types, lifting) cINF_greatest subset_empty) + with f show "\Y\common_refinement \

. X \ Y" + by (metis \\

\ {}\ cINF_greatest) +qed (use assms in auto) + +lemma finite_common_refinement: + assumes "finite \

" "\P. P \ \

\ finite P" + shows "finite (common_refinement \

)" +proof - + have "finite (\\<^sub>E P\\

. P)" + by (simp add: assms finite_PiE) + then show ?thesis + by (auto simp: common_refinement_def) +qed + +lemma card_common_refinement: + assumes "finite \

" "\P. P \ \

\ finite P" + shows "card (common_refinement \

) \ (\P \ \

. card P)" +proof - + have "card (common_refinement \

) \ card (\f \ (\\<^sub>E P\\

. P). {\ (f ` \

)})" + unfolding common_refinement_def by (meson card_Diff1_le) + also have "\ \ (\f\(\\<^sub>E P\\

. P). card{\ (f ` \

)})" + by (metis assms finite_PiE card_UN_le) + also have "\ = card(\\<^sub>E P\\

. P)" + by simp + also have "\ = (\P \ \

. card P)" + by (simp add: assms(1) card_PiE dual_order.eq_iff) + finally show ?thesis . +qed + end diff --git a/src/HOL/Set.thy b/src/HOL/Set.thy --- a/src/HOL/Set.thy +++ b/src/HOL/Set.thy @@ -1,2041 +1,2050 @@ (* Title: HOL/Set.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel *) section \Set theory for higher-order logic\ theory Set imports Lattices Boolean_Algebras begin subsection \Sets as predicates\ typedecl 'a set axiomatization Collect :: "('a \ bool) \ 'a set" \ \comprehension\ and member :: "'a \ 'a set \ bool" \ \membership\ where mem_Collect_eq [iff, code_unfold]: "member a (Collect P) = P a" and Collect_mem_eq [simp]: "Collect (\x. member x A) = A" notation member ("'(\')") and member ("(_/ \ _)" [51, 51] 50) abbreviation not_member where "not_member x A \ \ (x \ A)" \ \non-membership\ notation not_member ("'(\')") and not_member ("(_/ \ _)" [51, 51] 50) notation (ASCII) member ("'(:')") and member ("(_/ : _)" [51, 51] 50) and not_member ("'(~:')") and not_member ("(_/ ~: _)" [51, 51] 50) text \Set comprehensions\ syntax "_Coll" :: "pttrn \ bool \ 'a set" ("(1{_./ _})") translations "{x. P}" \ "CONST Collect (\x. P)" syntax (ASCII) "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{(_/: _)./ _})") syntax "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{(_/ \ _)./ _})") translations "{p:A. P}" \ "CONST Collect (\p. p \ A \ P)" lemma CollectI: "P a \ a \ {x. P x}" by simp lemma CollectD: "a \ {x. P x} \ P a" by simp lemma Collect_cong: "(\x. P x = Q x) \ {x. P x} = {x. Q x}" by simp text \ Simproc for pulling \x = t\ in \{x. \ \ x = t \ \}\ to the front (and similarly for \t = x\): \ simproc_setup defined_Collect ("{x. P x \ Q x}") = \ fn _ => Quantifier1.rearrange_Collect (fn ctxt => resolve_tac ctxt @{thms Collect_cong} 1 THEN resolve_tac ctxt @{thms iffI} 1 THEN ALLGOALS (EVERY' [REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}, DEPTH_SOLVE_1 o (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms conjI})])) \ lemmas CollectE = CollectD [elim_format] lemma set_eqI: assumes "\x. x \ A \ x \ B" shows "A = B" proof - from assms have "{x. x \ A} = {x. x \ B}" by simp then show ?thesis by simp qed lemma set_eq_iff: "A = B \ (\x. x \ A \ x \ B)" by (auto intro:set_eqI) lemma Collect_eqI: assumes "\x. P x = Q x" shows "Collect P = Collect Q" using assms by (auto intro: set_eqI) text \Lifting of predicate class instances\ instantiation set :: (type) boolean_algebra begin definition less_eq_set where "A \ B \ (\x. member x A) \ (\x. member x B)" definition less_set where "A < B \ (\x. member x A) < (\x. member x B)" definition inf_set where "A \ B = Collect ((\x. member x A) \ (\x. member x B))" definition sup_set where "A \ B = Collect ((\x. member x A) \ (\x. member x B))" definition bot_set where "\ = Collect \" definition top_set where "\ = Collect \" definition uminus_set where "- A = Collect (- (\x. member x A))" definition minus_set where "A - B = Collect ((\x. member x A) - (\x. member x B))" instance by standard (simp_all add: less_eq_set_def less_set_def inf_set_def sup_set_def bot_set_def top_set_def uminus_set_def minus_set_def less_le_not_le sup_inf_distrib1 diff_eq set_eqI fun_eq_iff del: inf_apply sup_apply bot_apply top_apply minus_apply uminus_apply) end text \Set enumerations\ abbreviation empty :: "'a set" ("{}") where "{} \ bot" definition insert :: "'a \ 'a set \ 'a set" where insert_compr: "insert a B = {x. x = a \ x \ B}" syntax "_Finset" :: "args \ 'a set" ("{(_)}") translations "{x, xs}" \ "CONST insert x {xs}" "{x}" \ "CONST insert x {}" subsection \Subsets and bounded quantifiers\ abbreviation subset :: "'a set \ 'a set \ bool" where "subset \ less" abbreviation subset_eq :: "'a set \ 'a set \ bool" where "subset_eq \ less_eq" notation subset ("'(\')") and subset ("(_/ \ _)" [51, 51] 50) and subset_eq ("'(\')") and subset_eq ("(_/ \ _)" [51, 51] 50) abbreviation (input) supset :: "'a set \ 'a set \ bool" where "supset \ greater" abbreviation (input) supset_eq :: "'a set \ 'a set \ bool" where "supset_eq \ greater_eq" notation supset ("'(\')") and supset ("(_/ \ _)" [51, 51] 50) and supset_eq ("'(\')") and supset_eq ("(_/ \ _)" [51, 51] 50) notation (ASCII output) subset ("'(<')") and subset ("(_/ < _)" [51, 51] 50) and subset_eq ("'(<=')") and subset_eq ("(_/ <= _)" [51, 51] 50) definition Ball :: "'a set \ ('a \ bool) \ bool" where "Ball A P \ (\x. x \ A \ P x)" \ \bounded universal quantifiers\ definition Bex :: "'a set \ ('a \ bool) \ bool" where "Bex A P \ (\x. x \ A \ P x)" \ \bounded existential quantifiers\ syntax (ASCII) "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3ALL (_/:_)./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3EX (_/:_)./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3EX! (_/:_)./ _)" [0, 0, 10] 10) "_Bleast" :: "id \ 'a set \ bool \ 'a" ("(3LEAST (_/:_)./ _)" [0, 0, 10] 10) syntax (input) "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3! (_/:_)./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3? (_/:_)./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3?! (_/:_)./ _)" [0, 0, 10] 10) syntax "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3\(_/\_)./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3\(_/\_)./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3\!(_/\_)./ _)" [0, 0, 10] 10) "_Bleast" :: "id \ 'a set \ bool \ 'a" ("(3LEAST(_/\_)./ _)" [0, 0, 10] 10) translations "\x\A. P" \ "CONST Ball A (\x. P)" "\x\A. P" \ "CONST Bex A (\x. P)" "\!x\A. P" \ "\!x. x \ A \ P" "LEAST x:A. P" \ "LEAST x. x \ A \ P" syntax (ASCII output) "_setlessAll" :: "[idt, 'a, bool] \ bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) "_setlessEx" :: "[idt, 'a, bool] \ bool" ("(3EX _<_./ _)" [0, 0, 10] 10) "_setleAll" :: "[idt, 'a, bool] \ bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) "_setleEx" :: "[idt, 'a, bool] \ bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) "_setleEx1" :: "[idt, 'a, bool] \ bool" ("(3EX! _<=_./ _)" [0, 0, 10] 10) syntax "_setlessAll" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setlessEx" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleAll" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleEx" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleEx1" :: "[idt, 'a, bool] \ bool" ("(3\!_\_./ _)" [0, 0, 10] 10) translations "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" "\!A\B. P" \ "\!A. A \ B \ P" print_translation \ let val All_binder = Mixfix.binder_name \<^const_syntax>\All\; val Ex_binder = Mixfix.binder_name \<^const_syntax>\Ex\; val impl = \<^const_syntax>\HOL.implies\; val conj = \<^const_syntax>\HOL.conj\; val sbset = \<^const_syntax>\subset\; val sbset_eq = \<^const_syntax>\subset_eq\; val trans = [((All_binder, impl, sbset), \<^syntax_const>\_setlessAll\), ((All_binder, impl, sbset_eq), \<^syntax_const>\_setleAll\), ((Ex_binder, conj, sbset), \<^syntax_const>\_setlessEx\), ((Ex_binder, conj, sbset_eq), \<^syntax_const>\_setleEx\)]; fun mk v (v', T) c n P = if v = v' andalso not (Term.exists_subterm (fn Free (x, _) => x = v | _ => false) n) then Syntax.const c $ Syntax_Trans.mark_bound_body (v', T) $ n $ P else raise Match; fun tr' q = (q, fn _ => (fn [Const (\<^syntax_const>\_bound\, _) $ Free (v, Type (\<^type_name>\set\, _)), Const (c, _) $ (Const (d, _) $ (Const (\<^syntax_const>\_bound\, _) $ Free (v', T)) $ n) $ P] => (case AList.lookup (=) trans (q, c, d) of NONE => raise Match | SOME l => mk v (v', T) l n P) | _ => raise Match)); in [tr' All_binder, tr' Ex_binder] end \ text \ \<^medskip> Translate between \{e | x1\xn. P}\ and \{u. \x1\xn. u = e \ P}\; \{y. \x1\xn. y = e \ P}\ is only translated if \[0..n] \ bvs e\. \ syntax "_Setcompr" :: "'a \ idts \ bool \ 'a set" ("(1{_ |/_./ _})") parse_translation \ let val ex_tr = snd (Syntax_Trans.mk_binder_tr ("EX ", \<^const_syntax>\Ex\)); fun nvars (Const (\<^syntax_const>\_idts\, _) $ _ $ idts) = nvars idts + 1 | nvars _ = 1; fun setcompr_tr ctxt [e, idts, b] = let val eq = Syntax.const \<^const_syntax>\HOL.eq\ $ Bound (nvars idts) $ e; val P = Syntax.const \<^const_syntax>\HOL.conj\ $ eq $ b; val exP = ex_tr ctxt [idts, P]; in Syntax.const \<^const_syntax>\Collect\ $ absdummy dummyT exP end; in [(\<^syntax_const>\_Setcompr\, setcompr_tr)] end \ print_translation \ [Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Ball\ \<^syntax_const>\_Ball\, Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Bex\ \<^syntax_const>\_Bex\] \ \ \to avoid eta-contraction of body\ print_translation \ let val ex_tr' = snd (Syntax_Trans.mk_binder_tr' (\<^const_syntax>\Ex\, "DUMMY")); fun setcompr_tr' ctxt [Abs (abs as (_, _, P))] = let fun check (Const (\<^const_syntax>\Ex\, _) $ Abs (_, _, P), n) = check (P, n + 1) | check (Const (\<^const_syntax>\HOL.conj\, _) $ (Const (\<^const_syntax>\HOL.eq\, _) $ Bound m $ e) $ P, n) = n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso subset (=) (0 upto (n - 1), add_loose_bnos (e, 0, [])) | check _ = false; fun tr' (_ $ abs) = let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr' ctxt [abs] in Syntax.const \<^syntax_const>\_Setcompr\ $ e $ idts $ Q end; in if check (P, 0) then tr' P else let val (x as _ $ Free(xN, _), t) = Syntax_Trans.atomic_abs_tr' abs; val M = Syntax.const \<^syntax_const>\_Coll\ $ x $ t; in case t of Const (\<^const_syntax>\HOL.conj\, _) $ (Const (\<^const_syntax>\Set.member\, _) $ (Const (\<^syntax_const>\_bound\, _) $ Free (yN, _)) $ A) $ P => if xN = yN then Syntax.const \<^syntax_const>\_Collect\ $ x $ A $ P else M | _ => M end end; in [(\<^const_syntax>\Collect\, setcompr_tr')] end \ simproc_setup defined_Bex ("\x\A. P x \ Q x") = \ fn _ => Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def}) \ simproc_setup defined_All ("\x\A. P x \ Q x") = \ fn _ => Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms Ball_def}) \ lemma ballI [intro!]: "(\x. x \ A \ P x) \ \x\A. P x" by (simp add: Ball_def) lemmas strip = impI allI ballI lemma bspec [dest?]: "\x\A. P x \ x \ A \ P x" by (simp add: Ball_def) text \Gives better instantiation for bound:\ setup \ map_theory_claset (fn ctxt => ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt')) \ ML \ structure Simpdata = struct open Simpdata; val mksimps_pairs = [(\<^const_name>\Ball\, @{thms bspec})] @ mksimps_pairs; end; open Simpdata; \ declaration \fn _ => Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))\ lemma ballE [elim]: "\x\A. P x \ (P x \ Q) \ (x \ A \ Q) \ Q" unfolding Ball_def by blast lemma bexI [intro]: "P x \ x \ A \ \x\A. P x" \ \Normally the best argument order: \P x\ constrains the choice of \x \ A\.\ unfolding Bex_def by blast lemma rev_bexI [intro?]: "x \ A \ P x \ \x\A. P x" \ \The best argument order when there is only one \x \ A\.\ unfolding Bex_def by blast lemma bexCI: "(\x\A. \ P x \ P a) \ a \ A \ \x\A. P x" unfolding Bex_def by blast lemma bexE [elim!]: "\x\A. P x \ (\x. x \ A \ P x \ Q) \ Q" unfolding Bex_def by blast lemma ball_triv [simp]: "(\x\A. P) \ ((\x. x \ A) \ P)" \ \trivial rewrite rule.\ by (simp add: Ball_def) lemma bex_triv [simp]: "(\x\A. P) \ ((\x. x \ A) \ P)" \ \Dual form for existentials.\ by (simp add: Bex_def) lemma bex_triv_one_point1 [simp]: "(\x\A. x = a) \ a \ A" by blast lemma bex_triv_one_point2 [simp]: "(\x\A. a = x) \ a \ A" by blast lemma bex_one_point1 [simp]: "(\x\A. x = a \ P x) \ a \ A \ P a" by blast lemma bex_one_point2 [simp]: "(\x\A. a = x \ P x) \ a \ A \ P a" by blast lemma ball_one_point1 [simp]: "(\x\A. x = a \ P x) \ (a \ A \ P a)" by blast lemma ball_one_point2 [simp]: "(\x\A. a = x \ P x) \ (a \ A \ P a)" by blast lemma ball_conj_distrib: "(\x\A. P x \ Q x) \ (\x\A. P x) \ (\x\A. Q x)" by blast lemma bex_disj_distrib: "(\x\A. P x \ Q x) \ (\x\A. P x) \ (\x\A. Q x)" by blast text \Congruence rules\ lemma ball_cong: "\ A = B; \x. x \ B \ P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: Ball_def) lemma ball_cong_simp [cong]: "\ A = B; \x. x \ B =simp=> P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Ball_def) lemma bex_cong: "\ A = B; \x. x \ B \ P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: Bex_def cong: conj_cong) lemma bex_cong_simp [cong]: "\ A = B; \x. x \ B =simp=> P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Bex_def cong: conj_cong) lemma bex1_def: "(\!x\X. P x) \ (\x\X. P x) \ (\x\X. \y\X. P x \ P y \ x = y)" by auto subsection \Basic operations\ subsubsection \Subsets\ lemma subsetI [intro!]: "(\x. x \ A \ x \ B) \ A \ B" by (simp add: less_eq_set_def le_fun_def) text \ \<^medskip> Map the type \'a set \ anything\ to just \'a\; for overloading constants whose first argument has type \'a set\. \ lemma subsetD [elim, intro?]: "A \ B \ c \ A \ c \ B" by (simp add: less_eq_set_def le_fun_def) \ \Rule in Modus Ponens style.\ lemma rev_subsetD [intro?,no_atp]: "c \ A \ A \ B \ c \ B" \ \The same, with reversed premises for use with @{method erule} -- cf. @{thm rev_mp}.\ by (rule subsetD) lemma subsetCE [elim,no_atp]: "A \ B \ (c \ A \ P) \ (c \ B \ P) \ P" \ \Classical elimination rule.\ by (auto simp add: less_eq_set_def le_fun_def) lemma subset_eq: "A \ B \ (\x\A. x \ B)" by blast lemma contra_subsetD [no_atp]: "A \ B \ c \ B \ c \ A" by blast lemma subset_refl: "A \ A" by (fact order_refl) (* already [iff] *) lemma subset_trans: "A \ B \ B \ C \ A \ C" by (fact order_trans) lemma subset_not_subset_eq [code]: "A \ B \ A \ B \ \ B \ A" by (fact less_le_not_le) lemma eq_mem_trans: "a = b \ b \ A \ a \ A" by simp lemmas basic_trans_rules [trans] = order_trans_rules rev_subsetD subsetD eq_mem_trans subsubsection \Equality\ lemma subset_antisym [intro!]: "A \ B \ B \ A \ A = B" \ \Anti-symmetry of the subset relation.\ by (iprover intro: set_eqI subsetD) text \\<^medskip> Equality rules from ZF set theory -- are they appropriate here?\ lemma equalityD1: "A = B \ A \ B" by simp lemma equalityD2: "A = B \ B \ A" by simp text \ \<^medskip> Be careful when adding this to the claset as \subset_empty\ is in the simpset: \<^prop>\A = {}\ goes to \<^prop>\{} \ A\ and \<^prop>\A \ {}\ and then back to \<^prop>\A = {}\! \ lemma equalityE: "A = B \ (A \ B \ B \ A \ P) \ P" by simp lemma equalityCE [elim]: "A = B \ (c \ A \ c \ B \ P) \ (c \ A \ c \ B \ P) \ P" by blast lemma eqset_imp_iff: "A = B \ x \ A \ x \ B" by simp lemma eqelem_imp_iff: "x = y \ x \ A \ y \ A" by simp subsubsection \The empty set\ lemma empty_def: "{} = {x. False}" by (simp add: bot_set_def bot_fun_def) lemma empty_iff [simp]: "c \ {} \ False" by (simp add: empty_def) lemma emptyE [elim!]: "a \ {} \ P" by simp lemma empty_subsetI [iff]: "{} \ A" \ \One effect is to delete the ASSUMPTION \<^prop>\{} \ A\\ by blast lemma equals0I: "(\y. y \ A \ False) \ A = {}" by blast lemma equals0D: "A = {} \ a \ A" \ \Use for reasoning about disjointness: \A \ B = {}\\ by blast lemma ball_empty [simp]: "Ball {} P \ True" by (simp add: Ball_def) lemma bex_empty [simp]: "Bex {} P \ False" by (simp add: Bex_def) subsubsection \The universal set -- UNIV\ abbreviation UNIV :: "'a set" where "UNIV \ top" lemma UNIV_def: "UNIV = {x. True}" by (simp add: top_set_def top_fun_def) lemma UNIV_I [simp]: "x \ UNIV" by (simp add: UNIV_def) declare UNIV_I [intro] \ \unsafe makes it less likely to cause problems\ lemma UNIV_witness [intro?]: "\x. x \ UNIV" by simp lemma subset_UNIV: "A \ UNIV" by (fact top_greatest) (* already simp *) text \ \<^medskip> Eta-contracting these two rules (to remove \P\) causes them to be ignored because of their interaction with congruence rules. \ lemma ball_UNIV [simp]: "Ball UNIV P \ All P" by (simp add: Ball_def) lemma bex_UNIV [simp]: "Bex UNIV P \ Ex P" by (simp add: Bex_def) lemma UNIV_eq_I: "(\x. x \ A) \ UNIV = A" by auto lemma UNIV_not_empty [iff]: "UNIV \ {}" by (blast elim: equalityE) lemma empty_not_UNIV[simp]: "{} \ UNIV" by blast subsubsection \The Powerset operator -- Pow\ definition Pow :: "'a set \ 'a set set" where Pow_def: "Pow A = {B. B \ A}" lemma Pow_iff [iff]: "A \ Pow B \ A \ B" by (simp add: Pow_def) lemma PowI: "A \ B \ A \ Pow B" by (simp add: Pow_def) lemma PowD: "A \ Pow B \ A \ B" by (simp add: Pow_def) lemma Pow_bottom: "{} \ Pow B" by simp lemma Pow_top: "A \ Pow A" by simp lemma Pow_not_empty: "Pow A \ {}" using Pow_top by blast subsubsection \Set complement\ lemma Compl_iff [simp]: "c \ - A \ c \ A" by (simp add: fun_Compl_def uminus_set_def) lemma ComplI [intro!]: "(c \ A \ False) \ c \ - A" by (simp add: fun_Compl_def uminus_set_def) blast text \ \<^medskip> This form, with negated conclusion, works well with the Classical prover. Negated assumptions behave like formulae on the right side of the notional turnstile \dots \ lemma ComplD [dest!]: "c \ - A \ c \ A" by simp lemmas ComplE = ComplD [elim_format] lemma Compl_eq: "- A = {x. \ x \ A}" by blast subsubsection \Binary intersection\ abbreviation inter :: "'a set \ 'a set \ 'a set" (infixl "\" 70) where "(\) \ inf" notation (ASCII) inter (infixl "Int" 70) lemma Int_def: "A \ B = {x. x \ A \ x \ B}" by (simp add: inf_set_def inf_fun_def) lemma Int_iff [simp]: "c \ A \ B \ c \ A \ c \ B" unfolding Int_def by blast lemma IntI [intro!]: "c \ A \ c \ B \ c \ A \ B" by simp lemma IntD1: "c \ A \ B \ c \ A" by simp lemma IntD2: "c \ A \ B \ c \ B" by simp lemma IntE [elim!]: "c \ A \ B \ (c \ A \ c \ B \ P) \ P" by simp lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" by (fact mono_inf) subsubsection \Binary union\ abbreviation union :: "'a set \ 'a set \ 'a set" (infixl "\" 65) where "union \ sup" notation (ASCII) union (infixl "Un" 65) lemma Un_def: "A \ B = {x. x \ A \ x \ B}" by (simp add: sup_set_def sup_fun_def) lemma Un_iff [simp]: "c \ A \ B \ c \ A \ c \ B" unfolding Un_def by blast lemma UnI1 [elim?]: "c \ A \ c \ A \ B" by simp lemma UnI2 [elim?]: "c \ B \ c \ A \ B" by simp text \\<^medskip> Classical introduction rule: no commitment to \A\ vs. \B\.\ lemma UnCI [intro!]: "(c \ B \ c \ A) \ c \ A \ B" by auto lemma UnE [elim!]: "c \ A \ B \ (c \ A \ P) \ (c \ B \ P) \ P" unfolding Un_def by blast lemma insert_def: "insert a B = {x. x = a} \ B" by (simp add: insert_compr Un_def) lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" by (fact mono_sup) subsubsection \Set difference\ lemma Diff_iff [simp]: "c \ A - B \ c \ A \ c \ B" by (simp add: minus_set_def fun_diff_def) lemma DiffI [intro!]: "c \ A \ c \ B \ c \ A - B" by simp lemma DiffD1: "c \ A - B \ c \ A" by simp lemma DiffD2: "c \ A - B \ c \ B \ P" by simp lemma DiffE [elim!]: "c \ A - B \ (c \ A \ c \ B \ P) \ P" by simp lemma set_diff_eq: "A - B = {x. x \ A \ x \ B}" by blast lemma Compl_eq_Diff_UNIV: "- A = (UNIV - A)" by blast subsubsection \Augmenting a set -- \<^const>\insert\\ lemma insert_iff [simp]: "a \ insert b A \ a = b \ a \ A" unfolding insert_def by blast lemma insertI1: "a \ insert a B" by simp lemma insertI2: "a \ B \ a \ insert b B" by simp lemma insertE [elim!]: "a \ insert b A \ (a = b \ P) \ (a \ A \ P) \ P" unfolding insert_def by blast lemma insertCI [intro!]: "(a \ B \ a = b) \ a \ insert b B" \ \Classical introduction rule.\ by auto lemma subset_insert_iff: "A \ insert x B \ (if x \ A then A - {x} \ B else A \ B)" by auto lemma set_insert: assumes "x \ A" obtains B where "A = insert x B" and "x \ B" proof show "A = insert x (A - {x})" using assms by blast show "x \ A - {x}" by blast qed lemma insert_ident: "x \ A \ x \ B \ insert x A = insert x B \ A = B" by auto lemma insert_eq_iff: assumes "a \ A" "b \ B" shows "insert a A = insert b B \ (if a = b then A = B else \C. A = insert b C \ b \ C \ B = insert a C \ a \ C)" (is "?L \ ?R") proof show ?R if ?L proof (cases "a = b") case True with assms \?L\ show ?R by (simp add: insert_ident) next case False let ?C = "A - {b}" have "A = insert b ?C \ b \ ?C \ B = insert a ?C \ a \ ?C" using assms \?L\ \a \ b\ by auto then show ?R using \a \ b\ by auto qed show ?L if ?R using that by (auto split: if_splits) qed lemma insert_UNIV: "insert x UNIV = UNIV" by auto subsubsection \Singletons, using insert\ lemma singletonI [intro!]: "a \ {a}" \ \Redundant? But unlike \insertCI\, it proves the subgoal immediately!\ by (rule insertI1) lemma singletonD [dest!]: "b \ {a} \ b = a" by blast lemmas singletonE = singletonD [elim_format] lemma singleton_iff: "b \ {a} \ b = a" by blast lemma singleton_inject [dest!]: "{a} = {b} \ a = b" by blast lemma singleton_insert_inj_eq [iff]: "{b} = insert a A \ a = b \ A \ {b}" by blast lemma singleton_insert_inj_eq' [iff]: "insert a A = {b} \ a = b \ A \ {b}" by blast lemma subset_singletonD: "A \ {x} \ A = {} \ A = {x}" by fast lemma subset_singleton_iff: "X \ {a} \ X = {} \ X = {a}" by blast lemma subset_singleton_iff_Uniq: "(\a. A \ {a}) \ (\\<^sub>\\<^sub>1x. x \ A)" unfolding Uniq_def by blast lemma singleton_conv [simp]: "{x. x = a} = {a}" by blast lemma singleton_conv2 [simp]: "{x. a = x} = {a}" by blast lemma Diff_single_insert: "A - {x} \ B \ A \ insert x B" by blast lemma subset_Diff_insert: "A \ B - insert x C \ A \ B - C \ x \ A" by blast lemma doubleton_eq_iff: "{a, b} = {c, d} \ a = c \ b = d \ a = d \ b = c" by (blast elim: equalityE) lemma Un_singleton_iff: "A \ B = {x} \ A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x}" by auto lemma singleton_Un_iff: "{x} = A \ B \ A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x}" by auto subsubsection \Image of a set under a function\ text \Frequently \b\ does not have the syntactic form of \f x\.\ definition image :: "('a \ 'b) \ 'a set \ 'b set" (infixr "`" 90) where "f ` A = {y. \x\A. y = f x}" lemma image_eqI [simp, intro]: "b = f x \ x \ A \ b \ f ` A" unfolding image_def by blast lemma imageI: "x \ A \ f x \ f ` A" by (rule image_eqI) (rule refl) lemma rev_image_eqI: "x \ A \ b = f x \ b \ f ` A" \ \This version's more effective when we already have the required \x\.\ by (rule image_eqI) lemma imageE [elim!]: assumes "b \ (\x. f x) ` A" \ \The eta-expansion gives variable-name preservation.\ obtains x where "b = f x" and "x \ A" using assms unfolding image_def by blast lemma Compr_image_eq: "{x \ f ` A. P x} = f ` {x \ A. P (f x)}" by auto lemma image_Un: "f ` (A \ B) = f ` A \ f ` B" by blast lemma image_iff: "z \ f ` A \ (\x\A. z = f x)" by blast lemma image_subsetI: "(\x. x \ A \ f x \ B) \ f ` A \ B" \ \Replaces the three steps \subsetI\, \imageE\, \hypsubst\, but breaks too many existing proofs.\ by blast lemma image_subset_iff: "f ` A \ B \ (\x\A. f x \ B)" \ \This rewrite rule would confuse users if made default.\ by blast lemma subset_imageE: assumes "B \ f ` A" obtains C where "C \ A" and "B = f ` C" proof - from assms have "B = f ` {a \ A. f a \ B}" by fast moreover have "{a \ A. f a \ B} \ A" by blast ultimately show thesis by (blast intro: that) qed lemma subset_image_iff: "B \ f ` A \ (\AA\A. B = f ` AA)" by (blast elim: subset_imageE) lemma image_ident [simp]: "(\x. x) ` Y = Y" by blast lemma image_empty [simp]: "f ` {} = {}" by blast lemma image_insert [simp]: "f ` insert a B = insert (f a) (f ` B)" by blast lemma image_constant: "x \ A \ (\x. c) ` A = {c}" by auto lemma image_constant_conv: "(\x. c) ` A = (if A = {} then {} else {c})" by auto lemma image_image: "f ` (g ` A) = (\x. f (g x)) ` A" by blast lemma insert_image [simp]: "x \ A \ insert (f x) (f ` A) = f ` A" by blast lemma image_is_empty [iff]: "f ` A = {} \ A = {}" by blast lemma empty_is_image [iff]: "{} = f ` A \ A = {}" by blast lemma image_Collect: "f ` {x. P x} = {f x | x. P x}" \ \NOT suitable as a default simp rule: the RHS isn't simpler than the LHS, with its implicit quantifier and conjunction. Also image enjoys better equational properties than does the RHS.\ by blast lemma if_image_distrib [simp]: "(\x. if P x then f x else g x) ` S = f ` (S \ {x. P x}) \ g ` (S \ {x. \ P x})" by auto lemma image_cong: "f ` M = g ` N" if "M = N" "\x. x \ N \ f x = g x" using that by (simp add: image_def) lemma image_cong_simp [cong]: "f ` M = g ` N" if "M = N" "\x. x \ N =simp=> f x = g x" using that image_cong [of M N f g] by (simp add: simp_implies_def) lemma image_Int_subset: "f ` (A \ B) \ f ` A \ f ` B" by blast lemma image_diff_subset: "f ` A - f ` B \ f ` (A - B)" by blast lemma Setcompr_eq_image: "{f x |x. x \ A} = f ` A" by blast lemma setcompr_eq_image: "{f x |x. P x} = f ` {x. P x}" by auto lemma ball_imageD: "\x\f ` A. P x \ \x\A. P (f x)" by simp lemma bex_imageD: "\x\f ` A. P x \ \x\A. P (f x)" by auto lemma image_add_0 [simp]: "(+) (0::'a::comm_monoid_add) ` S = S" by auto text \\<^medskip> Range of a function -- just an abbreviation for image!\ abbreviation range :: "('a \ 'b) \ 'b set" \ \of function\ where "range f \ f ` UNIV" lemma range_eqI: "b = f x \ b \ range f" by simp lemma rangeI: "f x \ range f" by simp lemma rangeE [elim?]: "b \ range (\x. f x) \ (\x. b = f x \ P) \ P" by (rule imageE) +lemma range_subsetD: "range f \ B \ f i \ B" + by blast + lemma full_SetCompr_eq: "{u. \x. u = f x} = range f" by auto lemma range_composition: "range (\x. f (g x)) = f ` range g" by auto lemma range_constant [simp]: "range (\_. x) = {x}" by (simp add: image_constant) lemma range_eq_singletonD: "range f = {a} \ f x = a" by auto subsubsection \Some rules with \if\\ text \Elimination of \{x. \ \ x = t \ \}\.\ lemma Collect_conv_if: "{x. x = a \ P x} = (if P a then {a} else {})" by auto lemma Collect_conv_if2: "{x. a = x \ P x} = (if P a then {a} else {})" by auto text \ Rewrite rules for boolean case-splitting: faster than \if_split [split]\. \ lemma if_split_eq1: "(if Q then x else y) = b \ (Q \ x = b) \ (\ Q \ y = b)" by (rule if_split) lemma if_split_eq2: "a = (if Q then x else y) \ (Q \ a = x) \ (\ Q \ a = y)" by (rule if_split) text \ Split ifs on either side of the membership relation. Not for \[simp]\ -- can cause goals to blow up! \ lemma if_split_mem1: "(if Q then x else y) \ b \ (Q \ x \ b) \ (\ Q \ y \ b)" by (rule if_split) lemma if_split_mem2: "(a \ (if Q then x else y)) \ (Q \ a \ x) \ (\ Q \ a \ y)" by (rule if_split [where P = "\S. a \ S"]) lemmas split_ifs = if_bool_eq_conj if_split_eq1 if_split_eq2 if_split_mem1 if_split_mem2 (*Would like to add these, but the existing code only searches for the outer-level constant, which in this case is just Set.member; we instead need to use term-nets to associate patterns with rules. Also, if a rule fails to apply, then the formula should be kept. [("uminus", Compl_iff RS iffD1), ("minus", [Diff_iff RS iffD1]), ("Int", [IntD1,IntD2]), ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] *) subsection \Further operations and lemmas\ subsubsection \The ``proper subset'' relation\ lemma psubsetI [intro!]: "A \ B \ A \ B \ A \ B" unfolding less_le by blast lemma psubsetE [elim!]: "A \ B \ (A \ B \ \ B \ A \ R) \ R" unfolding less_le by blast lemma psubset_insert_iff: "A \ insert x B \ (if x \ B then A \ B else if x \ A then A - {x} \ B else A \ B)" by (auto simp add: less_le subset_insert_iff) lemma psubset_eq: "A \ B \ A \ B \ A \ B" by (simp only: less_le) lemma psubset_imp_subset: "A \ B \ A \ B" by (simp add: psubset_eq) lemma psubset_trans: "A \ B \ B \ C \ A \ C" unfolding less_le by (auto dest: subset_antisym) lemma psubsetD: "A \ B \ c \ A \ c \ B" unfolding less_le by (auto dest: subsetD) lemma psubset_subset_trans: "A \ B \ B \ C \ A \ C" by (auto simp add: psubset_eq) lemma subset_psubset_trans: "A \ B \ B \ C \ A \ C" by (auto simp add: psubset_eq) lemma psubset_imp_ex_mem: "A \ B \ \b. b \ B - A" unfolding less_le by blast lemma atomize_ball: "(\x. x \ A \ P x) \ Trueprop (\x\A. P x)" by (simp only: Ball_def atomize_all atomize_imp) lemmas [symmetric, rulify] = atomize_ball and [symmetric, defn] = atomize_ball lemma image_Pow_mono: "f ` A \ B \ image f ` Pow A \ Pow B" by blast lemma image_Pow_surj: "f ` A = B \ image f ` Pow A = Pow B" by (blast elim: subset_imageE) subsubsection \Derived rules involving subsets.\ text \\insert\.\ lemma subset_insertI: "B \ insert a B" by (rule subsetI) (erule insertI2) lemma subset_insertI2: "A \ B \ A \ insert b B" by blast lemma subset_insert: "x \ A \ A \ insert x B \ A \ B" by blast text \\<^medskip> Finite Union -- the least upper bound of two sets.\ lemma Un_upper1: "A \ A \ B" by (fact sup_ge1) lemma Un_upper2: "B \ A \ B" by (fact sup_ge2) lemma Un_least: "A \ C \ B \ C \ A \ B \ C" by (fact sup_least) text \\<^medskip> Finite Intersection -- the greatest lower bound of two sets.\ lemma Int_lower1: "A \ B \ A" by (fact inf_le1) lemma Int_lower2: "A \ B \ B" by (fact inf_le2) lemma Int_greatest: "C \ A \ C \ B \ C \ A \ B" by (fact inf_greatest) text \\<^medskip> Set difference.\ lemma Diff_subset[simp]: "A - B \ A" by blast lemma Diff_subset_conv: "A - B \ C \ A \ B \ C" by blast subsubsection \Equalities involving union, intersection, inclusion, etc.\ text \\{}\.\ lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})" \ \supersedes \Collect_False_empty\\ by auto lemma subset_empty [simp]: "A \ {} \ A = {}" by (fact bot_unique) lemma not_psubset_empty [iff]: "\ (A < {})" by (fact not_less_bot) (* FIXME: already simp *) lemma Collect_subset [simp]: "{x\A. P x} \ A" by auto lemma Collect_empty_eq [simp]: "Collect P = {} \ (\x. \ P x)" by blast lemma empty_Collect_eq [simp]: "{} = Collect P \ (\x. \ P x)" by blast lemma Collect_neg_eq: "{x. \ P x} = - {x. P x}" by blast lemma Collect_disj_eq: "{x. P x \ Q x} = {x. P x} \ {x. Q x}" by blast lemma Collect_imp_eq: "{x. P x \ Q x} = - {x. P x} \ {x. Q x}" by blast lemma Collect_conj_eq: "{x. P x \ Q x} = {x. P x} \ {x. Q x}" by blast lemma Collect_mono_iff: "Collect P \ Collect Q \ (\x. P x \ Q x)" by blast text \\<^medskip> \insert\.\ lemma insert_is_Un: "insert a A = {a} \ A" \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a {}\\ by blast lemma insert_not_empty [simp]: "insert a A \ {}" and empty_not_insert [simp]: "{} \ insert a A" by blast+ lemma insert_absorb: "a \ A \ insert a A = A" \ \\[simp]\ causes recursive calls when there are nested inserts\ \ \with \<^emph>\quadratic\ running time\ by blast lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A" by blast lemma insert_commute: "insert x (insert y A) = insert y (insert x A)" by blast lemma insert_subset [simp]: "insert x A \ B \ x \ B \ A \ B" by blast lemma mk_disjoint_insert: "a \ A \ \B. A = insert a B \ a \ B" \ \use new \B\ rather than \A - {a}\ to avoid infinite unfolding\ by (rule exI [where x = "A - {a}"]) blast lemma insert_Collect: "insert a (Collect P) = {u. u \ a \ P u}" by auto lemma insert_inter_insert [simp]: "insert a A \ insert a B = insert a (A \ B)" by blast lemma insert_disjoint [simp]: "insert a A \ B = {} \ a \ B \ A \ B = {}" "{} = insert a A \ B \ a \ B \ {} = A \ B" by auto lemma disjoint_insert [simp]: "B \ insert a A = {} \ a \ B \ B \ A = {}" "{} = A \ insert b B \ b \ A \ {} = A \ B" by auto text \\<^medskip> \Int\\ lemma Int_absorb: "A \ A = A" by (fact inf_idem) (* already simp *) lemma Int_left_absorb: "A \ (A \ B) = A \ B" by (fact inf_left_idem) lemma Int_commute: "A \ B = B \ A" by (fact inf_commute) lemma Int_left_commute: "A \ (B \ C) = B \ (A \ C)" by (fact inf_left_commute) lemma Int_assoc: "(A \ B) \ C = A \ (B \ C)" by (fact inf_assoc) lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute \ \Intersection is an AC-operator\ lemma Int_absorb1: "B \ A \ A \ B = B" by (fact inf_absorb2) lemma Int_absorb2: "A \ B \ A \ B = A" by (fact inf_absorb1) lemma Int_empty_left: "{} \ B = {}" by (fact inf_bot_left) (* already simp *) lemma Int_empty_right: "A \ {} = {}" by (fact inf_bot_right) (* already simp *) lemma disjoint_eq_subset_Compl: "A \ B = {} \ A \ - B" by blast lemma disjoint_iff: "A \ B = {} \ (\x. x\A \ x \ B)" by blast lemma disjoint_iff_not_equal: "A \ B = {} \ (\x\A. \y\B. x \ y)" by blast lemma Int_UNIV_left: "UNIV \ B = B" by (fact inf_top_left) (* already simp *) lemma Int_UNIV_right: "A \ UNIV = A" by (fact inf_top_right) (* already simp *) lemma Int_Un_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by (fact inf_sup_distrib1) lemma Int_Un_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" by (fact inf_sup_distrib2) lemma Int_UNIV [simp]: "A \ B = UNIV \ A = UNIV \ B = UNIV" by (fact inf_eq_top_iff) (* already simp *) lemma Int_subset_iff [simp]: "C \ A \ B \ C \ A \ C \ B" by (fact le_inf_iff) lemma Int_Collect: "x \ A \ {x. P x} \ x \ A \ P x" by blast text \\<^medskip> \Un\.\ lemma Un_absorb: "A \ A = A" by (fact sup_idem) (* already simp *) lemma Un_left_absorb: "A \ (A \ B) = A \ B" by (fact sup_left_idem) lemma Un_commute: "A \ B = B \ A" by (fact sup_commute) lemma Un_left_commute: "A \ (B \ C) = B \ (A \ C)" by (fact sup_left_commute) lemma Un_assoc: "(A \ B) \ C = A \ (B \ C)" by (fact sup_assoc) lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute \ \Union is an AC-operator\ lemma Un_absorb1: "A \ B \ A \ B = B" by (fact sup_absorb2) lemma Un_absorb2: "B \ A \ A \ B = A" by (fact sup_absorb1) lemma Un_empty_left: "{} \ B = B" by (fact sup_bot_left) (* already simp *) lemma Un_empty_right: "A \ {} = A" by (fact sup_bot_right) (* already simp *) lemma Un_UNIV_left: "UNIV \ B = UNIV" by (fact sup_top_left) (* already simp *) lemma Un_UNIV_right: "A \ UNIV = UNIV" by (fact sup_top_right) (* already simp *) lemma Un_insert_left [simp]: "(insert a B) \ C = insert a (B \ C)" by blast lemma Un_insert_right [simp]: "A \ (insert a B) = insert a (A \ B)" by blast lemma Int_insert_left: "(insert a B) \ C = (if a \ C then insert a (B \ C) else B \ C)" by auto lemma Int_insert_left_if0 [simp]: "a \ C \ (insert a B) \ C = B \ C" by auto lemma Int_insert_left_if1 [simp]: "a \ C \ (insert a B) \ C = insert a (B \ C)" by auto lemma Int_insert_right: "A \ (insert a B) = (if a \ A then insert a (A \ B) else A \ B)" by auto lemma Int_insert_right_if0 [simp]: "a \ A \ A \ (insert a B) = A \ B" by auto lemma Int_insert_right_if1 [simp]: "a \ A \ A \ (insert a B) = insert a (A \ B)" by auto lemma Un_Int_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by (fact sup_inf_distrib1) lemma Un_Int_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" by (fact sup_inf_distrib2) lemma Un_Int_crazy: "(A \ B) \ (B \ C) \ (C \ A) = (A \ B) \ (B \ C) \ (C \ A)" by blast lemma subset_Un_eq: "A \ B \ A \ B = B" by (fact le_iff_sup) lemma Un_empty [iff]: "A \ B = {} \ A = {} \ B = {}" by (fact sup_eq_bot_iff) (* FIXME: already simp *) lemma Un_subset_iff [simp]: "A \ B \ C \ A \ C \ B \ C" by (fact le_sup_iff) lemma Un_Diff_Int: "(A - B) \ (A \ B) = A" by blast lemma Diff_Int2: "A \ C - B \ C = A \ C - B" by blast lemma subset_UnE: assumes "C \ A \ B" obtains A' B' where "A' \ A" "B' \ B" "C = A' \ B'" proof show "C \ A \ A" "C \ B \ B" "C = (C \ A) \ (C \ B)" using assms by blast+ qed lemma Un_Int_eq [simp]: "(S \ T) \ S = S" "(S \ T) \ T = T" "S \ (S \ T) = S" "T \ (S \ T) = T" by auto lemma Int_Un_eq [simp]: "(S \ T) \ S = S" "(S \ T) \ T = T" "S \ (S \ T) = S" "T \ (S \ T) = T" by auto text \\<^medskip> Set complement\ lemma Compl_disjoint [simp]: "A \ - A = {}" by (fact inf_compl_bot) lemma Compl_disjoint2 [simp]: "- A \ A = {}" by (fact compl_inf_bot) lemma Compl_partition: "A \ - A = UNIV" by (fact sup_compl_top) lemma Compl_partition2: "- A \ A = UNIV" by (fact compl_sup_top) lemma double_complement: "- (-A) = A" for A :: "'a set" by (fact double_compl) (* already simp *) lemma Compl_Un: "- (A \ B) = (- A) \ (- B)" by (fact compl_sup) (* already simp *) lemma Compl_Int: "- (A \ B) = (- A) \ (- B)" by (fact compl_inf) (* already simp *) lemma subset_Compl_self_eq: "A \ - A \ A = {}" by blast lemma Un_Int_assoc_eq: "(A \ B) \ C = A \ (B \ C) \ C \ A" \ \Halmos, Naive Set Theory, page 16.\ by blast lemma Compl_UNIV_eq: "- UNIV = {}" by (fact compl_top_eq) (* already simp *) lemma Compl_empty_eq: "- {} = UNIV" by (fact compl_bot_eq) (* already simp *) lemma Compl_subset_Compl_iff [iff]: "- A \ - B \ B \ A" by (fact compl_le_compl_iff) (* FIXME: already simp *) lemma Compl_eq_Compl_iff [iff]: "- A = - B \ A = B" for A B :: "'a set" by (fact compl_eq_compl_iff) (* FIXME: already simp *) lemma Compl_insert: "- insert x A = (- A) - {x}" by blast text \\<^medskip> Bounded quantifiers. The following are not added to the default simpset because (a) they duplicate the body and (b) there are no similar rules for \Int\. \ lemma ball_Un: "(\x \ A \ B. P x) \ (\x\A. P x) \ (\x\B. P x)" by blast lemma bex_Un: "(\x \ A \ B. P x) \ (\x\A. P x) \ (\x\B. P x)" by blast text \\<^medskip> Set difference.\ lemma Diff_eq: "A - B = A \ (- B)" by blast lemma Diff_eq_empty_iff [simp]: "A - B = {} \ A \ B" by blast lemma Diff_cancel [simp]: "A - A = {}" by blast lemma Diff_idemp [simp]: "(A - B) - B = A - B" for A B :: "'a set" by blast lemma Diff_triv: "A \ B = {} \ A - B = A" by (blast elim: equalityE) lemma empty_Diff [simp]: "{} - A = {}" by blast lemma Diff_empty [simp]: "A - {} = A" by blast lemma Diff_UNIV [simp]: "A - UNIV = {}" by blast lemma Diff_insert0 [simp]: "x \ A \ A - insert x B = A - B" by blast lemma Diff_insert: "A - insert a B = A - B - {a}" \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a 0\\ by blast lemma Diff_insert2: "A - insert a B = A - {a} - B" \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a 0\\ by blast lemma insert_Diff_if: "insert x A - B = (if x \ B then A - B else insert x (A - B))" by auto lemma insert_Diff1 [simp]: "x \ B \ insert x A - B = A - B" by blast lemma insert_Diff_single[simp]: "insert a (A - {a}) = insert a A" by blast lemma insert_Diff: "a \ A \ insert a (A - {a}) = A" by blast lemma Diff_insert_absorb: "x \ A \ (insert x A) - {x} = A" by auto lemma Diff_disjoint [simp]: "A \ (B - A) = {}" by blast lemma Diff_partition: "A \ B \ A \ (B - A) = B" by blast lemma double_diff: "A \ B \ B \ C \ B - (C - A) = A" by blast lemma Un_Diff_cancel [simp]: "A \ (B - A) = A \ B" by blast lemma Un_Diff_cancel2 [simp]: "(B - A) \ A = B \ A" by blast lemma Diff_Un: "A - (B \ C) = (A - B) \ (A - C)" by blast lemma Diff_Int: "A - (B \ C) = (A - B) \ (A - C)" by blast lemma Diff_Diff_Int: "A - (A - B) = A \ B" by blast lemma Un_Diff: "(A \ B) - C = (A - C) \ (B - C)" by blast lemma Int_Diff: "(A \ B) - C = A \ (B - C)" by blast lemma Diff_Int_distrib: "C \ (A - B) = (C \ A) - (C \ B)" by blast lemma Diff_Int_distrib2: "(A - B) \ C = (A \ C) - (B \ C)" by blast lemma Diff_Compl [simp]: "A - (- B) = A \ B" by auto lemma Compl_Diff_eq [simp]: "- (A - B) = - A \ B" by blast lemma subset_Compl_singleton [simp]: "A \ - {b} \ b \ A" by blast text \\<^medskip> Quantification over type \<^typ>\bool\.\ lemma bool_induct: "P True \ P False \ P x" by (cases x) auto lemma all_bool_eq: "(\b. P b) \ P True \ P False" by (auto intro: bool_induct) lemma bool_contrapos: "P x \ \ P False \ P True" by (cases x) auto lemma ex_bool_eq: "(\b. P b) \ P True \ P False" by (auto intro: bool_contrapos) lemma UNIV_bool: "UNIV = {False, True}" by (auto intro: bool_induct) text \\<^medskip> \Pow\\ lemma Pow_empty [simp]: "Pow {} = {{}}" by (auto simp add: Pow_def) lemma Pow_singleton_iff [simp]: "Pow X = {Y} \ X = {} \ Y = {}" by blast (* somewhat slow *) lemma Pow_insert: "Pow (insert a A) = Pow A \ (insert a ` Pow A)" by (blast intro: image_eqI [where ?x = "u - {a}" for u]) lemma Pow_Compl: "Pow (- A) = {- B | B. A \ Pow B}" by (blast intro: exI [where ?x = "- u" for u]) lemma Pow_UNIV [simp]: "Pow UNIV = UNIV" by blast lemma Un_Pow_subset: "Pow A \ Pow B \ Pow (A \ B)" by blast lemma Pow_Int_eq [simp]: "Pow (A \ B) = Pow A \ Pow B" by blast text \\<^medskip> Miscellany.\ +lemma Int_Diff_disjoint: "A \ B \ (A - B) = {}" + by blast + +lemma Int_Diff_Un: "A \ B \ (A - B) = A" + by blast + lemma set_eq_subset: "A = B \ A \ B \ B \ A" by blast lemma subset_iff: "A \ B \ (\t. t \ A \ t \ B)" by blast lemma subset_iff_psubset_eq: "A \ B \ A \ B \ A = B" unfolding less_le by blast lemma all_not_in_conv [simp]: "(\x. x \ A) \ A = {}" by blast lemma ex_in_conv: "(\x. x \ A) \ A \ {}" by blast lemma ball_simps [simp, no_atp]: "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" "\P. (\x\{}. P x) \ True" "\P. (\x\UNIV. P x) \ (\x. P x)" "\a B P. (\x\insert a B. P x) \ (P a \ (\x\B. P x))" "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" "\A P. (\ (\x\A. P x)) \ (\x\A. \ P x)" by auto lemma bex_simps [simp, no_atp]: "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" "\P. (\x\{}. P x) \ False" "\P. (\x\UNIV. P x) \ (\x. P x)" "\a B P. (\x\insert a B. P x) \ (P a \ (\x\B. P x))" "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" "\A P. (\(\x\A. P x)) \ (\x\A. \ P x)" by auto lemma ex_image_cong_iff [simp, no_atp]: "(\x. x\f`A) \ A \ {}" "(\x. x\f`A \ P x) \ (\x\A. P (f x))" by auto subsubsection \Monotonicity of various operations\ lemma image_mono: "A \ B \ f ` A \ f ` B" by blast lemma Pow_mono: "A \ B \ Pow A \ Pow B" by blast lemma insert_mono: "C \ D \ insert a C \ insert a D" by blast lemma Un_mono: "A \ C \ B \ D \ A \ B \ C \ D" by (fact sup_mono) lemma Int_mono: "A \ C \ B \ D \ A \ B \ C \ D" by (fact inf_mono) lemma Diff_mono: "A \ C \ D \ B \ A - B \ C - D" by blast lemma Compl_anti_mono: "A \ B \ - B \ - A" by (fact compl_mono) text \\<^medskip> Monotonicity of implications.\ lemma in_mono: "A \ B \ x \ A \ x \ B" by (rule impI) (erule subsetD) lemma conj_mono: "P1 \ Q1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover lemma disj_mono: "P1 \ Q1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover lemma imp_mono: "Q1 \ P1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover lemma imp_refl: "P \ P" .. lemma not_mono: "Q \ P \ \ P \ \ Q" by iprover lemma ex_mono: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover lemma all_mono: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover lemma Collect_mono: "(\x. P x \ Q x) \ Collect P \ Collect Q" by blast lemma Int_Collect_mono: "A \ B \ (\x. x \ A \ P x \ Q x) \ A \ Collect P \ B \ Collect Q" by blast lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono Collect_mono in_mono lemma eq_to_mono: "a = b \ c = d \ b \ d \ a \ c" by iprover subsubsection \Inverse image of a function\ definition vimage :: "('a \ 'b) \ 'b set \ 'a set" (infixr "-`" 90) where "f -` B \ {x. f x \ B}" lemma vimage_eq [simp]: "a \ f -` B \ f a \ B" unfolding vimage_def by blast lemma vimage_singleton_eq: "a \ f -` {b} \ f a = b" by simp lemma vimageI [intro]: "f a = b \ b \ B \ a \ f -` B" unfolding vimage_def by blast lemma vimageI2: "f a \ A \ a \ f -` A" unfolding vimage_def by fast lemma vimageE [elim!]: "a \ f -` B \ (\x. f a = x \ x \ B \ P) \ P" unfolding vimage_def by blast lemma vimageD: "a \ f -` A \ f a \ A" unfolding vimage_def by fast lemma vimage_empty [simp]: "f -` {} = {}" by blast lemma vimage_Compl: "f -` (- A) = - (f -` A)" by blast lemma vimage_Un [simp]: "f -` (A \ B) = (f -` A) \ (f -` B)" by blast lemma vimage_Int [simp]: "f -` (A \ B) = (f -` A) \ (f -` B)" by fast lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}" by blast lemma vimage_Collect: "(\x. P (f x) = Q x) \ f -` (Collect P) = Collect Q" by blast lemma vimage_insert: "f -` (insert a B) = (f -` {a}) \ (f -` B)" \ \NOT suitable for rewriting because of the recurrence of \{a}\.\ by blast lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)" by blast lemma vimage_UNIV [simp]: "f -` UNIV = UNIV" by blast lemma vimage_mono: "A \ B \ f -` A \ f -` B" \ \monotonicity\ by blast lemma vimage_image_eq: "f -` (f ` A) = {y. \x\A. f x = f y}" by (blast intro: sym) lemma image_vimage_subset: "f ` (f -` A) \ A" by blast lemma image_vimage_eq [simp]: "f ` (f -` A) = A \ range f" by blast lemma image_subset_iff_subset_vimage: "f ` A \ B \ A \ f -` B" by blast lemma subset_vimage_iff: "A \ f -` B \ (\x\A. f x \ B)" by auto lemma vimage_const [simp]: "((\x. c) -` A) = (if c \ A then UNIV else {})" by auto lemma vimage_if [simp]: "((\x. if x \ B then c else d) -` A) = (if c \ A then (if d \ A then UNIV else B) else if d \ A then - B else {})" by (auto simp add: vimage_def) lemma vimage_inter_cong: "(\ w. w \ S \ f w = g w) \ f -` y \ S = g -` y \ S" by auto lemma vimage_ident [simp]: "(\x. x) -` Y = Y" by blast subsubsection \Singleton sets\ definition is_singleton :: "'a set \ bool" where "is_singleton A \ (\x. A = {x})" lemma is_singletonI [simp, intro!]: "is_singleton {x}" unfolding is_singleton_def by simp lemma is_singletonI': "A \ {} \ (\x y. x \ A \ y \ A \ x = y) \ is_singleton A" unfolding is_singleton_def by blast lemma is_singletonE: "is_singleton A \ (\x. A = {x} \ P) \ P" unfolding is_singleton_def by blast subsubsection \Getting the contents of a singleton set\ definition the_elem :: "'a set \ 'a" where "the_elem X = (THE x. X = {x})" lemma the_elem_eq [simp]: "the_elem {x} = x" by (simp add: the_elem_def) lemma is_singleton_the_elem: "is_singleton A \ A = {the_elem A}" by (auto simp: is_singleton_def) lemma the_elem_image_unique: assumes "A \ {}" and *: "\y. y \ A \ f y = f x" shows "the_elem (f ` A) = f x" unfolding the_elem_def proof (rule the1_equality) from \A \ {}\ obtain y where "y \ A" by auto with * have "f x = f y" by simp with \y \ A\ have "f x \ f ` A" by blast with * show "f ` A = {f x}" by auto then show "\!x. f ` A = {x}" by auto qed subsubsection \Least value operator\ lemma Least_mono: "mono f \ \x\S. \y\S. x \ y \ (LEAST y. y \ f ` S) = f (LEAST x. x \ S)" for f :: "'a::order \ 'b::order" \ \Courtesy of Stephan Merz\ apply clarify apply (erule_tac P = "\x. x \ S" in LeastI2_order) apply fast apply (rule LeastI2_order) apply (auto elim: monoD intro!: order_antisym) done subsubsection \Monad operation\ definition bind :: "'a set \ ('a \ 'b set) \ 'b set" where "bind A f = {x. \B \ f`A. x \ B}" hide_const (open) bind lemma bind_bind: "Set.bind (Set.bind A B) C = Set.bind A (\x. Set.bind (B x) C)" for A :: "'a set" by (auto simp: bind_def) lemma empty_bind [simp]: "Set.bind {} f = {}" by (simp add: bind_def) lemma nonempty_bind_const: "A \ {} \ Set.bind A (\_. B) = B" by (auto simp: bind_def) lemma bind_const: "Set.bind A (\_. B) = (if A = {} then {} else B)" by (auto simp: bind_def) lemma bind_singleton_conv_image: "Set.bind A (\x. {f x}) = f ` A" by (auto simp: bind_def) subsubsection \Operations for execution\ definition is_empty :: "'a set \ bool" where [code_abbrev]: "is_empty A \ A = {}" hide_const (open) is_empty definition remove :: "'a \ 'a set \ 'a set" where [code_abbrev]: "remove x A = A - {x}" hide_const (open) remove lemma member_remove [simp]: "x \ Set.remove y A \ x \ A \ x \ y" by (simp add: remove_def) definition filter :: "('a \ bool) \ 'a set \ 'a set" where [code_abbrev]: "filter P A = {a \ A. P a}" hide_const (open) filter lemma member_filter [simp]: "x \ Set.filter P A \ x \ A \ P x" by (simp add: filter_def) instantiation set :: (equal) equal begin definition "HOL.equal A B \ A \ B \ B \ A" instance by standard (auto simp add: equal_set_def) end text \Misc\ definition pairwise :: "('a \ 'a \ bool) \ 'a set \ bool" where "pairwise R S \ (\x \ S. \y \ S. x \ y \ R x y)" lemma pairwise_alt: "pairwise R S \ (\x\S. \y\S-{x}. R x y)" by (auto simp add: pairwise_def) lemma pairwise_trivial [simp]: "pairwise (\i j. j \ i) I" by (auto simp: pairwise_def) lemma pairwiseI [intro?]: "pairwise R S" if "\x y. x \ S \ y \ S \ x \ y \ R x y" using that by (simp add: pairwise_def) lemma pairwiseD: "R x y" and "R y x" if "pairwise R S" "x \ S" and "y \ S" and "x \ y" using that by (simp_all add: pairwise_def) lemma pairwise_empty [simp]: "pairwise P {}" by (simp add: pairwise_def) lemma pairwise_singleton [simp]: "pairwise P {A}" by (simp add: pairwise_def) lemma pairwise_insert: "pairwise r (insert x s) \ (\y. y \ s \ y \ x \ r x y \ r y x) \ pairwise r s" by (force simp: pairwise_def) lemma pairwise_subset: "pairwise P S \ T \ S \ pairwise P T" by (force simp: pairwise_def) lemma pairwise_mono: "\pairwise P A; \x y. P x y \ Q x y; B \ A\ \ pairwise Q B" by (fastforce simp: pairwise_def) lemma pairwise_imageI: "pairwise P (f ` A)" if "\x y. x \ A \ y \ A \ x \ y \ f x \ f y \ P (f x) (f y)" using that by (auto intro: pairwiseI) lemma pairwise_image: "pairwise r (f ` s) \ pairwise (\x y. (f x \ f y) \ r (f x) (f y)) s" by (force simp: pairwise_def) definition disjnt :: "'a set \ 'a set \ bool" where "disjnt A B \ A \ B = {}" lemma disjnt_self_iff_empty [simp]: "disjnt S S \ S = {}" by (auto simp: disjnt_def) lemma disjnt_iff: "disjnt A B \ (\x. \ (x \ A \ x \ B))" by (force simp: disjnt_def) lemma disjnt_sym: "disjnt A B \ disjnt B A" using disjnt_iff by blast lemma disjnt_empty1 [simp]: "disjnt {} A" and disjnt_empty2 [simp]: "disjnt A {}" by (auto simp: disjnt_def) lemma disjnt_insert1 [simp]: "disjnt (insert a X) Y \ a \ Y \ disjnt X Y" by (simp add: disjnt_def) lemma disjnt_insert2 [simp]: "disjnt Y (insert a X) \ a \ Y \ disjnt Y X" by (simp add: disjnt_def) lemma disjnt_subset1 : "\disjnt X Y; Z \ X\ \ disjnt Z Y" by (auto simp: disjnt_def) lemma disjnt_subset2 : "\disjnt X Y; Z \ Y\ \ disjnt X Z" by (auto simp: disjnt_def) lemma disjnt_Un1 [simp]: "disjnt (A \ B) C \ disjnt A C \ disjnt B C" by (auto simp: disjnt_def) lemma disjnt_Un2 [simp]: "disjnt C (A \ B) \ disjnt C A \ disjnt C B" by (auto simp: disjnt_def) lemma disjoint_image_subset: "\pairwise disjnt \; \X. X \ \ \ f X \ X\ \ pairwise disjnt (f `\)" unfolding disjnt_def pairwise_def by fast lemma pairwise_disjnt_iff: "pairwise disjnt \ \ (\x. \\<^sub>\\<^sub>1 X. X \ \ \ x \ X)" by (auto simp: Uniq_def disjnt_iff pairwise_def) lemma disjnt_insert: \<^marker>\contributor \Lars Hupel\\ \disjnt (insert x M) N\ if \x \ N\ \disjnt M N\ using that by (simp add: disjnt_def) lemma Int_emptyI: "(\x. x \ A \ x \ B \ False) \ A \ B = {}" by blast lemma in_image_insert_iff: assumes "\C. C \ B \ x \ C" shows "A \ insert x ` B \ x \ A \ A - {x} \ B" (is "?P \ ?Q") proof assume ?P then show ?Q using assms by auto next assume ?Q then have "x \ A" and "A - {x} \ B" by simp_all from \A - {x} \ B\ have "insert x (A - {x}) \ insert x ` B" by (rule imageI) also from \x \ A\ have "insert x (A - {x}) = A" by auto finally show ?P . qed hide_const (open) member not_member lemmas equalityI = subset_antisym lemmas set_mp = subsetD lemmas set_rev_mp = rev_subsetD ML \ val Ball_def = @{thm Ball_def} val Bex_def = @{thm Bex_def} val CollectD = @{thm CollectD} val CollectE = @{thm CollectE} val CollectI = @{thm CollectI} val Collect_conj_eq = @{thm Collect_conj_eq} val Collect_mem_eq = @{thm Collect_mem_eq} val IntD1 = @{thm IntD1} val IntD2 = @{thm IntD2} val IntE = @{thm IntE} val IntI = @{thm IntI} val Int_Collect = @{thm Int_Collect} val UNIV_I = @{thm UNIV_I} val UNIV_witness = @{thm UNIV_witness} val UnE = @{thm UnE} val UnI1 = @{thm UnI1} val UnI2 = @{thm UnI2} val ballE = @{thm ballE} val ballI = @{thm ballI} val bexCI = @{thm bexCI} val bexE = @{thm bexE} val bexI = @{thm bexI} val bex_triv = @{thm bex_triv} val bspec = @{thm bspec} val contra_subsetD = @{thm contra_subsetD} val equalityCE = @{thm equalityCE} val equalityD1 = @{thm equalityD1} val equalityD2 = @{thm equalityD2} val equalityE = @{thm equalityE} val equalityI = @{thm equalityI} val imageE = @{thm imageE} val imageI = @{thm imageI} val image_Un = @{thm image_Un} val image_insert = @{thm image_insert} val insert_commute = @{thm insert_commute} val insert_iff = @{thm insert_iff} val mem_Collect_eq = @{thm mem_Collect_eq} val rangeE = @{thm rangeE} val rangeI = @{thm rangeI} val range_eqI = @{thm range_eqI} val subsetCE = @{thm subsetCE} val subsetD = @{thm subsetD} val subsetI = @{thm subsetI} val subset_refl = @{thm subset_refl} val subset_trans = @{thm subset_trans} val vimageD = @{thm vimageD} val vimageE = @{thm vimageE} val vimageI = @{thm vimageI} val vimageI2 = @{thm vimageI2} val vimage_Collect = @{thm vimage_Collect} val vimage_Int = @{thm vimage_Int} val vimage_Un = @{thm vimage_Un} \ end