diff -r 8e0eece7058d -r 856c68ab6f13 src/HOL/Equiv_Relations.thy --- a/src/HOL/Equiv_Relations.thy Fri Mar 27 22:06:46 2020 +0100 +++ b/src/HOL/Equiv_Relations.thy Sat Mar 28 17:27:01 2020 +0000 @@ -34,21 +34,21 @@ unfolding refl_on_def by blast lemma equiv_comp_eq: "equiv A r \ r\ O r = r" - apply (unfold equiv_def) - apply clarify - apply (rule equalityI) - apply (iprover intro: sym_trans_comp_subset refl_on_comp_subset)+ - done + unfolding equiv_def + by (iprover intro: sym_trans_comp_subset refl_on_comp_subset equalityI) text \Second half.\ -lemma comp_equivI: "r\ O r = r \ Domain r = A \ equiv A r" - apply (unfold equiv_def refl_on_def sym_def trans_def) - apply (erule equalityE) - apply (subgoal_tac "\x y. (x, y) \ r \ (y, x) \ r") - apply fast - apply fast - done +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\ @@ -58,10 +58,7 @@ unfolding equiv_def trans_def sym_def by blast theorem equiv_class_eq: "equiv A r \ (a, b) \ r \ r``{a} = r``{b}" - apply (assumption | rule equalityI equiv_class_subset)+ - apply (unfold equiv_def sym_def) - apply blast - done + 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 @@ -91,7 +88,7 @@ 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" +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" @@ -101,32 +98,31 @@ 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 = {}" - apply (unfold quotient_def) - apply clarify - apply (rule equiv_class_eq) - apply assumption - apply (unfold equiv_def trans_def sym_def) - apply blast - done + unfolding quotient_def equiv_def trans_def sym_def by blast lemma quotient_eqI: - "equiv A r \ X \ A//r \ Y \ A//r \ x \ X \ y \ Y \ (x, y) \ r \ X = Y" - apply (clarify elim!: quotientE) - apply (rule equiv_class_eq) - apply assumption - apply (unfold equiv_def sym_def trans_def) - apply blast - done + 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: - "equiv A r \ X \ A//r \ Y \ A//r \ x \ X \ y \ Y \ X = Y \ (x, y) \ r" - apply (rule iffI) - prefer 2 - apply (blast del: equalityI intro: quotient_eqI) - apply (clarify elim!: quotientE) - apply (unfold equiv_def sym_def trans_def) - apply blast - done + 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) @@ -189,22 +185,22 @@ \ \lemma required to prove \UN_equiv_class\\ by auto -lemma UN_equiv_class: "equiv A r \ f respects r \ a \ A \ (\x \ r``{a}. f x) = f a" +lemma UN_equiv_class: + assumes "equiv A r" "f respects r" "a \ A" + shows "(\x \ r``{a}. f x) = f a" \ \Conversion rule\ - apply (rule equiv_class_self [THEN UN_constant_eq]) - apply assumption - apply assumption - apply (unfold equiv_def congruent_def sym_def) - apply (blast del: equalityI) - done +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: - "equiv A r \ f respects r \ X \ A//r \ (\x. x \ A \ f x \ B) \ (\x \ X. f x) \ B" - apply (unfold quotient_def) - apply clarify - apply (subst UN_equiv_class) - apply auto - done + 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! @@ -213,19 +209,23 @@ \ lemma UN_equiv_class_inject: - "equiv A r \ f respects r \ - (\x \ X. f x) = (\y \ Y. f y) \ X \ A//r ==> Y \ A//r - \ (\x y. x \ A \ y \ A \ f x = f y \ (x, y) \ r) - \ X = Y" - apply (unfold quotient_def) - apply clarify - apply (rule equiv_class_eq) - apply assumption - apply (subgoal_tac "f x = f xa") - apply blast - apply (erule box_equals) - apply (assumption | rule UN_equiv_class)+ - done + 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\ @@ -253,15 +253,20 @@ unfolding congruent_def congruent2_def equiv_def refl_on_def by blast lemma congruent2_implies_congruent_UN: - "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ a \ A2 \ - congruent r1 (\x1. \x2 \ r2``{a}. f x1 x2)" - apply (unfold congruent_def) - apply clarify - apply (rule equiv_type [THEN subsetD, THEN SigmaE2], assumption+) - apply (simp add: UN_equiv_class congruent2_implies_congruent) - apply (unfold congruent2_def equiv_def refl_on_def) - apply (blast del: equalityI) - done + 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 \ @@ -273,11 +278,10 @@ \ X1 \ A1//r1 \ X2 \ A2//r2 \ (\x1 x2. x1 \ A1 \ x2 \ A2 \ f x1 x2 \ B) \ (\x1 \ X1. \x2 \ X2. f x1 x2) \ B" - apply (unfold quotient_def) - apply clarify - apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN - congruent2_implies_congruent quotientI) - done + 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) = @@ -293,60 +297,63 @@ \ congruent2 r1 r2 f" \ \Suggested by John Harrison -- the two subproofs may be\ \ \\<^emph>\much\ simpler than the direct proof.\ - apply (unfold congruent2_def equiv_def refl_on_def) - apply clarify - apply (blast intro: trans) - done + 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" - apply (rule congruent2I [OF equivA equivA]) - apply (rule commute [THEN trans]) - apply (rule_tac [3] commute [THEN trans, symmetric]) - apply (rule_tac [5] sym) - apply (rule congt | assumption | - erule equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2])+ - done +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: "finite A \ r \ A \ A \ finite (A//r)" - \ \recall @{thm equiv_type}\ - apply (rule finite_subset) - apply (erule_tac [2] finite_Pow_iff [THEN iffD2]) - apply (unfold quotient_def) - apply blast - done +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" - apply (unfold quotient_def) - apply (rule finite_subset) - prefer 2 apply assumption - apply blast - done + unfolding quotient_def + by (erule rev_finite_subset) blast -lemma equiv_imp_dvd_card: "finite A \ equiv A r \ \X \ A//r. k dvd card X \ k dvd card A" - apply (rule Union_quotient [THEN subst [where P="\A. k dvd card A"]]) - apply assumption - apply (rule dvd_partition) - prefer 3 apply (blast dest: quotient_disj) - apply (simp_all add: Union_quotient equiv_type) - done +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: "finite A \ inj_on (\x. {x} // r) A \ card (A//r) = card A" - apply (simp add:quotient_def) - apply (subst card_UN_disjoint) - apply assumption - apply simp - apply (fastforce simp add:inj_on_def) - apply simp - done +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 subsection \Projection\