haftmann@29655: (* Authors: Lawrence C Paulson, Cambridge University Computer Laboratory paulson@15300: Copyright 1996 University of Cambridge paulson@15300: *) paulson@15300: paulson@15300: header {* Equivalence Relations in Higher-Order Set Theory *} paulson@15300: paulson@15300: theory Equiv_Relations haftmann@51112: imports Big_Operators Relation paulson@15300: begin paulson@15300: haftmann@40812: subsection {* Equivalence relations -- set version *} paulson@15300: haftmann@40812: definition equiv :: "'a set \ ('a \ 'a) set \ bool" where haftmann@40812: "equiv A r \ refl_on A r \ sym r \ trans r" paulson@15300: haftmann@40815: lemma equivI: haftmann@40815: "refl_on A r \ sym r \ trans r \ equiv A r" haftmann@40815: by (simp add: equiv_def) haftmann@40815: haftmann@40815: lemma equivE: haftmann@40815: assumes "equiv A r" haftmann@40815: obtains "refl_on A r" and "sym r" and "trans r" haftmann@40815: using assms by (simp add: equiv_def) haftmann@40815: paulson@15300: text {* paulson@15300: Suppes, Theorem 70: @{text r} is an equiv relation iff @{text "r\ O paulson@15300: r = r"}. paulson@15300: paulson@15300: First half: @{text "equiv A r ==> r\ O r = r"}. paulson@15300: *} paulson@15300: paulson@15300: lemma sym_trans_comp_subset: paulson@15300: "sym r ==> trans r ==> r\ O r \ r" haftmann@46752: by (unfold trans_def sym_def converse_unfold) blast paulson@15300: nipkow@30198: lemma refl_on_comp_subset: "refl_on A r ==> r \ r\ O r" nipkow@30198: by (unfold refl_on_def) blast paulson@15300: paulson@15300: lemma equiv_comp_eq: "equiv A r ==> r\ O r = r" paulson@15300: apply (unfold equiv_def) paulson@15300: apply clarify paulson@15300: apply (rule equalityI) nipkow@30198: apply (iprover intro: sym_trans_comp_subset refl_on_comp_subset)+ paulson@15300: done paulson@15300: paulson@15300: text {* Second half. *} paulson@15300: paulson@15300: lemma comp_equivI: paulson@15300: "r\ O r = r ==> Domain r = A ==> equiv A r" nipkow@30198: apply (unfold equiv_def refl_on_def sym_def trans_def) paulson@15300: apply (erule equalityE) paulson@15300: apply (subgoal_tac "\x y. (x, y) \ r --> (y, x) \ r") paulson@15300: apply fast paulson@15300: apply fast paulson@15300: done paulson@15300: paulson@15300: paulson@15300: subsection {* Equivalence classes *} paulson@15300: paulson@15300: lemma equiv_class_subset: paulson@15300: "equiv A r ==> (a, b) \ r ==> r``{a} \ r``{b}" paulson@15300: -- {* lemma for the next result *} paulson@15300: by (unfold equiv_def trans_def sym_def) blast paulson@15300: paulson@15300: theorem equiv_class_eq: "equiv A r ==> (a, b) \ r ==> r``{a} = r``{b}" paulson@15300: apply (assumption | rule equalityI equiv_class_subset)+ paulson@15300: apply (unfold equiv_def sym_def) paulson@15300: apply blast paulson@15300: done paulson@15300: paulson@15300: lemma equiv_class_self: "equiv A r ==> a \ A ==> a \ r``{a}" nipkow@30198: by (unfold equiv_def refl_on_def) blast paulson@15300: paulson@15300: lemma subset_equiv_class: paulson@15300: "equiv A r ==> r``{b} \ r``{a} ==> b \ A ==> (a,b) \ r" paulson@15300: -- {* lemma for the next result *} nipkow@30198: by (unfold equiv_def refl_on_def) blast paulson@15300: paulson@15300: lemma eq_equiv_class: paulson@15300: "r``{a} = r``{b} ==> equiv A r ==> b \ A ==> (a, b) \ r" nipkow@17589: by (iprover intro: equalityD2 subset_equiv_class) paulson@15300: paulson@15300: lemma equiv_class_nondisjoint: paulson@15300: "equiv A r ==> x \ (r``{a} \ r``{b}) ==> (a, b) \ r" paulson@15300: by (unfold equiv_def trans_def sym_def) blast paulson@15300: paulson@15300: lemma equiv_type: "equiv A r ==> r \ A \ A" nipkow@30198: by (unfold equiv_def refl_on_def) blast paulson@15300: paulson@15300: theorem equiv_class_eq_iff: paulson@15300: "equiv A r ==> ((x, y) \ r) = (r``{x} = r``{y} & x \ A & y \ A)" paulson@15300: by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) paulson@15300: paulson@15300: theorem eq_equiv_class_iff: paulson@15300: "equiv A r ==> x \ A ==> y \ A ==> (r``{x} = r``{y}) = ((x, y) \ r)" paulson@15300: by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) paulson@15300: paulson@15300: paulson@15300: subsection {* Quotients *} paulson@15300: haftmann@28229: definition quotient :: "'a set \ ('a \ 'a) set \ 'a set set" (infixl "'/'/" 90) where haftmann@37767: "A//r = (\x \ A. {r``{x}})" -- {* set of equiv classes *} paulson@15300: paulson@15300: lemma quotientI: "x \ A ==> r``{x} \ A//r" paulson@15300: by (unfold quotient_def) blast paulson@15300: paulson@15300: lemma quotientE: paulson@15300: "X \ A//r ==> (!!x. X = r``{x} ==> x \ A ==> P) ==> P" paulson@15300: by (unfold quotient_def) blast paulson@15300: paulson@15300: lemma Union_quotient: "equiv A r ==> Union (A//r) = A" nipkow@30198: by (unfold equiv_def refl_on_def quotient_def) blast paulson@15300: paulson@15300: lemma quotient_disj: paulson@15300: "equiv A r ==> X \ A//r ==> Y \ A//r ==> X = Y | (X \ Y = {})" paulson@15300: apply (unfold quotient_def) paulson@15300: apply clarify paulson@15300: apply (rule equiv_class_eq) paulson@15300: apply assumption paulson@15300: apply (unfold equiv_def trans_def sym_def) paulson@15300: apply blast paulson@15300: done paulson@15300: paulson@15300: lemma quotient_eqI: paulson@15300: "[|equiv A r; X \ A//r; Y \ A//r; x \ X; y \ Y; (x,y) \ r|] ==> X = Y" paulson@15300: apply (clarify elim!: quotientE) paulson@15300: apply (rule equiv_class_eq, assumption) paulson@15300: apply (unfold equiv_def sym_def trans_def, blast) paulson@15300: done paulson@15300: paulson@15300: lemma quotient_eq_iff: paulson@15300: "[|equiv A r; X \ A//r; Y \ A//r; x \ X; y \ Y|] ==> (X = Y) = ((x,y) \ r)" paulson@15300: apply (rule iffI) paulson@15300: prefer 2 apply (blast del: equalityI intro: quotient_eqI) paulson@15300: apply (clarify elim!: quotientE) paulson@15300: apply (unfold equiv_def sym_def trans_def, blast) paulson@15300: done paulson@15300: nipkow@18493: lemma eq_equiv_class_iff2: nipkow@18493: "\ equiv A r; x \ A; y \ A \ \ ({x}//r = {y}//r) = ((x,y) : r)" nipkow@18493: by(simp add:quotient_def eq_equiv_class_iff) nipkow@18493: paulson@15300: paulson@15300: lemma quotient_empty [simp]: "{}//r = {}" paulson@15300: by(simp add: quotient_def) paulson@15300: paulson@15300: lemma quotient_is_empty [iff]: "(A//r = {}) = (A = {})" paulson@15300: by(simp add: quotient_def) paulson@15300: paulson@15300: lemma quotient_is_empty2 [iff]: "({} = A//r) = (A = {})" paulson@15300: by(simp add: quotient_def) paulson@15300: paulson@15300: nipkow@15302: lemma singleton_quotient: "{x}//r = {r `` {x}}" nipkow@15302: by(simp add:quotient_def) nipkow@15302: nipkow@15302: lemma quotient_diff1: nipkow@15302: "\ inj_on (%a. {a}//r) A; a \ A \ \ (A - {a})//r = A//r - {a}//r" nipkow@15302: apply(simp add:quotient_def inj_on_def) nipkow@15302: apply blast nipkow@15302: done nipkow@15302: paulson@15300: subsection {* Defining unary operations upon equivalence classes *} paulson@15300: paulson@15300: text{*A congruence-preserving function*} haftmann@40816: haftmann@44278: definition congruent :: "('a \ 'a) set \ ('a \ 'b) \ bool" where haftmann@40817: "congruent r f \ (\(y, z) \ r. f y = f z)" haftmann@40816: haftmann@40816: lemma congruentI: haftmann@40816: "(\y z. (y, z) \ r \ f y = f z) \ congruent r f" haftmann@40817: by (auto simp add: congruent_def) haftmann@40816: haftmann@40816: lemma congruentD: haftmann@40816: "congruent r f \ (y, z) \ r \ f y = f z" haftmann@40817: by (auto simp add: congruent_def) paulson@15300: wenzelm@19363: abbreviation wenzelm@21404: RESPECTS :: "('a => 'b) => ('a * 'a) set => bool" wenzelm@21404: (infixr "respects" 80) where wenzelm@19363: "f respects r == congruent r f" paulson@15300: paulson@15300: paulson@15300: lemma UN_constant_eq: "a \ A ==> \y \ A. f y = c ==> (\y \ A. f(y))=c" paulson@15300: -- {* lemma required to prove @{text UN_equiv_class} *} paulson@15300: by auto paulson@15300: paulson@15300: lemma UN_equiv_class: paulson@15300: "equiv A r ==> f respects r ==> a \ A paulson@15300: ==> (\x \ r``{a}. f x) = f a" paulson@15300: -- {* Conversion rule *} paulson@15300: apply (rule equiv_class_self [THEN UN_constant_eq], assumption+) paulson@15300: apply (unfold equiv_def congruent_def sym_def) paulson@15300: apply (blast del: equalityI) paulson@15300: done paulson@15300: paulson@15300: lemma UN_equiv_class_type: paulson@15300: "equiv A r ==> f respects r ==> X \ A//r ==> paulson@15300: (!!x. x \ A ==> f x \ B) ==> (\x \ X. f x) \ B" paulson@15300: apply (unfold quotient_def) paulson@15300: apply clarify paulson@15300: apply (subst UN_equiv_class) paulson@15300: apply auto paulson@15300: done paulson@15300: paulson@15300: text {* paulson@15300: Sufficient conditions for injectiveness. Could weaken premises! paulson@15300: major premise could be an inclusion; bcong could be @{text "!!y. y \ paulson@15300: A ==> f y \ B"}. paulson@15300: *} paulson@15300: paulson@15300: lemma UN_equiv_class_inject: paulson@15300: "equiv A r ==> f respects r ==> paulson@15300: (\x \ X. f x) = (\y \ Y. f y) ==> X \ A//r ==> Y \ A//r paulson@15300: ==> (!!x y. x \ A ==> y \ A ==> f x = f y ==> (x, y) \ r) paulson@15300: ==> X = Y" paulson@15300: apply (unfold quotient_def) paulson@15300: apply clarify paulson@15300: apply (rule equiv_class_eq) paulson@15300: apply assumption paulson@15300: apply (subgoal_tac "f x = f xa") paulson@15300: apply blast paulson@15300: apply (erule box_equals) paulson@15300: apply (assumption | rule UN_equiv_class)+ paulson@15300: done paulson@15300: paulson@15300: paulson@15300: subsection {* Defining binary operations upon equivalence classes *} paulson@15300: paulson@15300: text{*A congruence-preserving function of two arguments*} haftmann@40817: haftmann@44278: definition congruent2 :: "('a \ 'a) set \ ('b \ 'b) set \ ('a \ 'b \ 'c) \ bool" where haftmann@40817: "congruent2 r1 r2 f \ (\(y1, z1) \ r1. \(y2, z2) \ r2. f y1 y2 = f z1 z2)" haftmann@40817: haftmann@40817: lemma congruent2I': haftmann@40817: assumes "\y1 z1 y2 z2. (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" haftmann@40817: shows "congruent2 r1 r2 f" haftmann@40817: using assms by (auto simp add: congruent2_def) haftmann@40817: haftmann@40817: lemma congruent2D: haftmann@40817: "congruent2 r1 r2 f \ (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" haftmann@40817: using assms by (auto simp add: congruent2_def) paulson@15300: paulson@15300: text{*Abbreviation for the common case where the relations are identical*} nipkow@19979: abbreviation wenzelm@21404: RESPECTS2:: "['a => 'a => 'b, ('a * 'a) set] => bool" wenzelm@21749: (infixr "respects2" 80) where nipkow@19979: "f respects2 r == congruent2 r r f" nipkow@19979: paulson@15300: paulson@15300: lemma congruent2_implies_congruent: paulson@15300: "equiv A r1 ==> congruent2 r1 r2 f ==> a \ A ==> congruent r2 (f a)" nipkow@30198: by (unfold congruent_def congruent2_def equiv_def refl_on_def) blast paulson@15300: paulson@15300: lemma congruent2_implies_congruent_UN: paulson@15300: "equiv A1 r1 ==> equiv A2 r2 ==> congruent2 r1 r2 f ==> a \ A2 ==> paulson@15300: congruent r1 (\x1. \x2 \ r2``{a}. f x1 x2)" paulson@15300: apply (unfold congruent_def) paulson@15300: apply clarify paulson@15300: apply (rule equiv_type [THEN subsetD, THEN SigmaE2], assumption+) paulson@15300: apply (simp add: UN_equiv_class congruent2_implies_congruent) nipkow@30198: apply (unfold congruent2_def equiv_def refl_on_def) paulson@15300: apply (blast del: equalityI) paulson@15300: done paulson@15300: paulson@15300: lemma UN_equiv_class2: paulson@15300: "equiv A1 r1 ==> equiv A2 r2 ==> congruent2 r1 r2 f ==> a1 \ A1 ==> a2 \ A2 paulson@15300: ==> (\x1 \ r1``{a1}. \x2 \ r2``{a2}. f x1 x2) = f a1 a2" paulson@15300: by (simp add: UN_equiv_class congruent2_implies_congruent paulson@15300: congruent2_implies_congruent_UN) paulson@15300: paulson@15300: lemma UN_equiv_class_type2: paulson@15300: "equiv A1 r1 ==> equiv A2 r2 ==> congruent2 r1 r2 f paulson@15300: ==> X1 \ A1//r1 ==> X2 \ A2//r2 paulson@15300: ==> (!!x1 x2. x1 \ A1 ==> x2 \ A2 ==> f x1 x2 \ B) paulson@15300: ==> (\x1 \ X1. \x2 \ X2. f x1 x2) \ B" paulson@15300: apply (unfold quotient_def) paulson@15300: apply clarify paulson@15300: apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN paulson@15300: congruent2_implies_congruent quotientI) paulson@15300: done paulson@15300: paulson@15300: lemma UN_UN_split_split_eq: paulson@15300: "(\(x1, x2) \ X. \(y1, y2) \ Y. A x1 x2 y1 y2) = paulson@15300: (\x \ X. \y \ Y. (\(x1, x2). (\(y1, y2). A x1 x2 y1 y2) y) x)" paulson@15300: -- {* Allows a natural expression of binary operators, *} paulson@15300: -- {* without explicit calls to @{text split} *} paulson@15300: by auto paulson@15300: paulson@15300: lemma congruent2I: paulson@15300: "equiv A1 r1 ==> equiv A2 r2 paulson@15300: ==> (!!y z w. w \ A2 ==> (y,z) \ r1 ==> f y w = f z w) paulson@15300: ==> (!!y z w. w \ A1 ==> (y,z) \ r2 ==> f w y = f w z) paulson@15300: ==> congruent2 r1 r2 f" paulson@15300: -- {* Suggested by John Harrison -- the two subproofs may be *} paulson@15300: -- {* \emph{much} simpler than the direct proof. *} nipkow@30198: apply (unfold congruent2_def equiv_def refl_on_def) paulson@15300: apply clarify paulson@15300: apply (blast intro: trans) paulson@15300: done paulson@15300: paulson@15300: lemma congruent2_commuteI: paulson@15300: assumes equivA: "equiv A r" paulson@15300: and commute: "!!y z. y \ A ==> z \ A ==> f y z = f z y" paulson@15300: and congt: "!!y z w. w \ A ==> (y,z) \ r ==> f w y = f w z" paulson@15300: shows "f respects2 r" paulson@15300: apply (rule congruent2I [OF equivA equivA]) paulson@15300: apply (rule commute [THEN trans]) paulson@15300: apply (rule_tac [3] commute [THEN trans, symmetric]) paulson@15300: apply (rule_tac [5] sym) haftmann@25482: apply (rule congt | assumption | paulson@15300: erule equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2])+ paulson@15300: done paulson@15300: haftmann@24728: haftmann@24728: subsection {* Quotients and finiteness *} haftmann@24728: wenzelm@40945: text {*Suggested by Florian Kammüller*} haftmann@24728: haftmann@24728: lemma finite_quotient: "finite A ==> r \ A \ A ==> finite (A//r)" haftmann@24728: -- {* recall @{thm equiv_type} *} haftmann@24728: apply (rule finite_subset) haftmann@24728: apply (erule_tac [2] finite_Pow_iff [THEN iffD2]) haftmann@24728: apply (unfold quotient_def) haftmann@24728: apply blast haftmann@24728: done haftmann@24728: haftmann@24728: lemma finite_equiv_class: haftmann@24728: "finite A ==> r \ A \ A ==> X \ A//r ==> finite X" haftmann@24728: apply (unfold quotient_def) haftmann@24728: apply (rule finite_subset) haftmann@24728: prefer 2 apply assumption haftmann@24728: apply blast haftmann@24728: done haftmann@24728: haftmann@24728: lemma equiv_imp_dvd_card: haftmann@24728: "finite A ==> equiv A r ==> \X \ A//r. k dvd card X haftmann@24728: ==> k dvd card A" berghofe@26791: apply (rule Union_quotient [THEN subst [where P="\A. k dvd card A"]]) haftmann@24728: apply assumption haftmann@24728: apply (rule dvd_partition) haftmann@24728: prefer 3 apply (blast dest: quotient_disj) haftmann@24728: apply (simp_all add: Union_quotient equiv_type) haftmann@24728: done haftmann@24728: haftmann@24728: lemma card_quotient_disjoint: haftmann@24728: "\ finite A; inj_on (\x. {x} // r) A \ \ card(A//r) = card A" haftmann@24728: apply(simp add:quotient_def) haftmann@24728: apply(subst card_UN_disjoint) haftmann@24728: apply assumption haftmann@24728: apply simp nipkow@44890: apply(fastforce simp add:inj_on_def) huffman@35216: apply simp haftmann@24728: done haftmann@24728: haftmann@40812: haftmann@40812: subsection {* Equivalence relations -- predicate version *} haftmann@40812: haftmann@40812: text {* Partial equivalences *} haftmann@40812: haftmann@40812: definition part_equivp :: "('a \ 'a \ bool) \ bool" where haftmann@40812: "part_equivp R \ (\x. R x x) \ (\x y. R x y \ R x x \ R y y \ R x = R y)" haftmann@40812: -- {* John-Harrison-style characterization *} haftmann@40812: haftmann@40812: lemma part_equivpI: haftmann@40812: "(\x. R x x) \ symp R \ transp R \ part_equivp R" haftmann@45969: by (auto simp add: part_equivp_def) (auto elim: sympE transpE) haftmann@40812: haftmann@40812: lemma part_equivpE: haftmann@40812: assumes "part_equivp R" haftmann@40812: obtains x where "R x x" and "symp R" and "transp R" haftmann@40812: proof - haftmann@40812: from assms have 1: "\x. R x x" haftmann@40812: and 2: "\x y. R x y \ R x x \ R y y \ R x = R y" haftmann@40812: by (unfold part_equivp_def) blast+ haftmann@40812: from 1 obtain x where "R x x" .. haftmann@40812: moreover have "symp R" haftmann@40812: proof (rule sympI) haftmann@40812: fix x y haftmann@40812: assume "R x y" haftmann@40812: with 2 [of x y] show "R y x" by auto haftmann@40812: qed haftmann@40812: moreover have "transp R" haftmann@40812: proof (rule transpI) haftmann@40812: fix x y z haftmann@40812: assume "R x y" and "R y z" haftmann@40812: with 2 [of x y] 2 [of y z] show "R x z" by auto haftmann@40812: qed haftmann@40812: ultimately show thesis by (rule that) haftmann@40812: qed haftmann@40812: haftmann@40812: lemma part_equivp_refl_symp_transp: haftmann@40812: "part_equivp R \ (\x. R x x) \ symp R \ transp R" haftmann@40812: by (auto intro: part_equivpI elim: part_equivpE) haftmann@40812: haftmann@40812: lemma part_equivp_symp: haftmann@40812: "part_equivp R \ R x y \ R y x" haftmann@40812: by (erule part_equivpE, erule sympE) haftmann@40812: haftmann@40812: lemma part_equivp_transp: haftmann@40812: "part_equivp R \ R x y \ R y z \ R x z" haftmann@40812: by (erule part_equivpE, erule transpE) haftmann@40812: haftmann@40812: lemma part_equivp_typedef: kaliszyk@44204: "part_equivp R \ \d. d \ {c. \x. R x x \ c = Collect (R x)}" kaliszyk@44204: by (auto elim: part_equivpE) haftmann@40812: haftmann@40812: haftmann@40812: text {* Total equivalences *} haftmann@40812: haftmann@40812: definition equivp :: "('a \ 'a \ bool) \ bool" where haftmann@40812: "equivp R \ (\x y. R x y = (R x = R y))" -- {* John-Harrison-style characterization *} haftmann@40812: haftmann@40812: lemma equivpI: haftmann@40812: "reflp R \ symp R \ transp R \ equivp R" haftmann@45969: by (auto elim: reflpE sympE transpE simp add: equivp_def) haftmann@40812: haftmann@40812: lemma equivpE: haftmann@40812: assumes "equivp R" haftmann@40812: obtains "reflp R" and "symp R" and "transp R" haftmann@40812: using assms by (auto intro!: that reflpI sympI transpI simp add: equivp_def) haftmann@40812: haftmann@40812: lemma equivp_implies_part_equivp: haftmann@40812: "equivp R \ part_equivp R" haftmann@40812: by (auto intro: part_equivpI elim: equivpE reflpE) haftmann@40812: haftmann@40812: lemma equivp_equiv: haftmann@40812: "equiv UNIV A \ equivp (\x y. (x, y) \ A)" haftmann@46752: by (auto intro!: equivI equivpI [to_set] elim!: equivE equivpE [to_set]) haftmann@40812: haftmann@40812: lemma equivp_reflp_symp_transp: haftmann@40812: shows "equivp R \ reflp R \ symp R \ transp R" haftmann@40812: by (auto intro: equivpI elim: equivpE) haftmann@40812: haftmann@40812: lemma identity_equivp: haftmann@40812: "equivp (op =)" haftmann@40812: by (auto intro: equivpI reflpI sympI transpI) haftmann@40812: haftmann@40812: lemma equivp_reflp: haftmann@40812: "equivp R \ R x x" haftmann@40812: by (erule equivpE, erule reflpE) haftmann@40812: haftmann@40812: lemma equivp_symp: haftmann@40812: "equivp R \ R x y \ R y x" haftmann@40812: by (erule equivpE, erule sympE) haftmann@40812: haftmann@40812: lemma equivp_transp: haftmann@40812: "equivp R \ R x y \ R y z \ R x z" haftmann@40812: by (erule equivpE, erule transpE) haftmann@40812: paulson@15300: end haftmann@46752: