# HG changeset patch # User wenzelm # Date 1470859500 -7200 # Node ID 4453cfb745e5b4c66d0c7ab2e26fcb0871f7aef9 # Parent 804b80a80016578e1f3a3a09e1ac5b70f1862de3 misc tuning and modernization; diff -r 804b80a80016 -r 4453cfb745e5 src/HOL/Equiv_Relations.thy --- a/src/HOL/Equiv_Relations.thy Wed Aug 10 22:03:58 2016 +0200 +++ b/src/HOL/Equiv_Relations.thy Wed Aug 10 22:05:00 2016 +0200 @@ -1,20 +1,19 @@ -(* Authors: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1996 University of Cambridge +(* 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 Relation + imports Groups_Big Relation 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" +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" +lemma equivI: "refl_on A r \ sym r \ trans r \ equiv A r" by (simp add: equiv_def) lemma equivE: @@ -23,20 +22,18 @@ using assms by (simp add: equiv_def) text \ - Suppes, Theorem 70: \r\ is an equiv relation iff \r\ O - r = r\. + Suppes, Theorem 70: \r\ is an equiv relation iff \r\ O r = r\. - First half: \equiv A r ==> 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" - by (unfold trans_def sym_def converse_unfold) blast +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" - by (unfold refl_on_def) 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" +lemma equiv_comp_eq: "equiv A r \ r\ O r = r" apply (unfold equiv_def) apply clarify apply (rule equalityI) @@ -45,11 +42,10 @@ text \Second half.\ -lemma comp_equivI: - "r\ O r = r ==> Domain r = A ==> equiv A r" +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 (subgoal_tac "\x y. (x, y) \ r \ (y, x) \ r") apply fast apply fast done @@ -57,62 +53,54 @@ subsection \Equivalence classes\ -lemma equiv_class_subset: - "equiv A r ==> (a, b) \ r ==> r``{a} \ r``{b}" +lemma equiv_class_subset: "equiv A r \ (a, b) \ r \ r``{a} \ r``{b}" \ \lemma for the next result\ - by (unfold equiv_def trans_def sym_def) blast + unfolding equiv_def trans_def sym_def by blast -theorem equiv_class_eq: "equiv A r ==> (a, b) \ r ==> r``{a} = r``{b}" +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 -lemma equiv_class_self: "equiv A r ==> a \ A ==> a \ r``{a}" - by (unfold equiv_def refl_on_def) blast +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 subset_equiv_class: "equiv A r \ r``{b} \ r``{a} \ b \ A \ (a, b) \ r" \ \lemma for the next result\ - by (unfold equiv_def refl_on_def) blast + unfolding equiv_def refl_on_def by blast -lemma eq_equiv_class: - "r``{a} = r``{b} ==> equiv A r ==> b \ A ==> (a, b) \ r" +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" - by (unfold equiv_def trans_def sym_def) blast +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" - by (unfold equiv_def refl_on_def) blast +lemma equiv_type: "equiv A r \ r \ A \ A" + unfolding equiv_def refl_on_def by blast -theorem equiv_class_eq_iff: - "equiv A r ==> ((x, y) \ r) = (r``{x} = r``{y} & x \ A & y \ A)" +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) -theorem eq_equiv_class_iff: - "equiv A r ==> x \ A ==> y \ A ==> (r``{x} = r``{y}) = ((x, y) \ r)" +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) 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\ +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" - by (unfold quotient_def) blast + unfolding quotient_def by blast -lemma quotientE: - "X \ A//r ==> (!!x. X = r``{x} ==> x \ A ==> P) ==> P" - by (unfold quotient_def) 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" - by (unfold equiv_def refl_on_def quotient_def) 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 = {})" +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) @@ -122,108 +110,96 @@ done lemma quotient_eqI: - "[|equiv A r; X \ A//r; Y \ A//r; x \ X; y \ Y; (x,y) \ r|] ==> X = Y" + "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, assumption) - apply (unfold equiv_def sym_def trans_def, blast) + apply (rule equiv_class_eq) + apply assumption + apply (unfold equiv_def sym_def trans_def) + apply blast done 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) + "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, blast) + apply (unfold equiv_def sym_def trans_def) + apply blast done -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 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) + by (simp add: quotient_def) -lemma quotient_is_empty [iff]: "(A//r = {}) = (A = {})" -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 quotient_is_empty2 [iff]: "{} = A//r \ A = {}" + by (simp add: quotient_def) lemma singleton_quotient: "{x}//r = {r `` {x}}" -by(simp add:quotient_def) + by (simp add: quotient_def) -lemma quotient_diff1: - "\ inj_on (%a. {a}//r) A; a \ A \ \ (A - {a})//r = A//r - {a}//r" -apply(simp add:quotient_def inj_on_def) -apply blast -done +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}" +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}" +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" +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)" - apply (erule finite_surj [where f = "\X. S``X"]) - apply (simp add: refines_equiv_image_eq) - done + "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)" - apply (subst refines_equiv_image_eq [of R S A, symmetric]) - apply (auto simp: card_image_le [where f = "\X. S``X"]) - done + "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\ +text \A congruence-preserving function.\ -definition congruent :: "('a \ 'a) set \ ('a \ 'b) \ bool" where - "congruent r f \ (\(y, z) \ r. f y = f z)" +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" +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" +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" +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 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: - "equiv A r ==> f respects r ==> a \ A - ==> (\x \ r``{a}. f x) = f a" +lemma UN_equiv_class: "equiv A r \ f respects r \ a \ A \ (\x \ r``{a}. f x) = f a" \ \Conversion rule\ - apply (rule equiv_class_self [THEN UN_constant_eq], assumption+) + 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 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" + "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) @@ -232,15 +208,15 @@ text \ Sufficient conditions for injectiveness. Could weaken premises! - major premise could be an inclusion; bcong could be \!!y. y \ - A ==> f y \ B\. + major premise could be an inclusion; \bcong\ could be + \\y. y \ A \ f y \ B\. \ 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" + "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) @@ -254,33 +230,30 @@ subsection \Defining binary operations upon equivalence classes\ -text\A congruence-preserving function of two arguments\ +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)" +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" +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" +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)" - by (unfold congruent_def congruent2_def equiv_def refl_on_def) blast + "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: - "equiv A1 r1 ==> equiv A2 r2 ==> congruent2 r1 r2 f ==> a \ A2 ==> + "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 @@ -291,20 +264,19 @@ done 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) + "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" + "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" apply (unfold quotient_def) apply clarify apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN - congruent2_implies_congruent quotientI) + congruent2_implies_congruent quotientI) done lemma UN_UN_split_split_eq: @@ -315,12 +287,12 @@ 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" + "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.\ + \ \\<^emph>\much\ simpler than the direct proof.\ apply (unfold congruent2_def equiv_def refl_on_def) apply clarify apply (blast intro: trans) @@ -328,8 +300,8 @@ 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" + 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]) @@ -344,7 +316,7 @@ text \Suggested by Florian Kammüller\ -lemma finite_quotient: "finite A ==> r \ A \ A ==> finite (A//r)" +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]) @@ -352,93 +324,94 @@ apply blast done -lemma finite_equiv_class: - "finite A ==> r \ A \ A ==> X \ A//r ==> finite X" +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 -lemma equiv_imp_dvd_card: - "finite A ==> equiv A r ==> \X \ A//r. k dvd card X - ==> k dvd card A" +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) + prefer 3 apply (blast dest: quotient_disj) + apply (simp_all add: Union_quotient equiv_type) done -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 +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 - apply(fastforce simp add:inj_on_def) -apply simp -done + done subsection \Projection\ -definition proj where "proj r x = r `` {x}" +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_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)" -apply(rule iffI, auto simp add: proj_preserves) -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 - thus "x \ A" using assms unfolding equiv_def refl_on_def by blast + 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 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 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_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_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_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 +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\ +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)" +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" +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: @@ -447,7 +420,7 @@ 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" - by (unfold part_equivp_def) blast+ + unfolding part_equivp_def by blast+ from 1 obtain x where "R x x" .. moreover have "symp R" proof (rule sympI) @@ -464,30 +437,25 @@ ultimately show thesis by (rule that) qed -lemma part_equivp_refl_symp_transp: - "part_equivp R \ (\x. R x x) \ symp R \ transp R" +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" +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" +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)}" +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\ +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\ +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" +lemma equivpI: "reflp R \ symp R \ transp R \ equivp R" by (auto elim: reflpE sympE transpE simp add: equivp_def) lemma equivpE: @@ -495,32 +463,25 @@ 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" +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)" +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: - shows "equivp R \ reflp R \ symp R \ transp R" +lemma equivp_reflp_symp_transp: "equivp R \ reflp R \ symp R \ transp R" by (auto intro: equivpI elim: equivpE) -lemma identity_equivp: - "equivp (op =)" +lemma identity_equivp: "equivp (op =)" by (auto intro: equivpI reflpI sympI transpI) -lemma equivp_reflp: - "equivp R \ R x x" +lemma equivp_reflp: "equivp R \ R x x" by (erule equivpE, erule reflpE) -lemma equivp_symp: - "equivp R \ R x y \ R y x" +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" +lemma equivp_transp: "equivp R \ R x y \ R y z \ R x z" by (erule equivpE, erule transpE) hide_const (open) proj