diff -r 33a08cfc3ae5 -r 2b5da07a0b89 src/HOL/Integ/Equiv.thy --- a/src/HOL/Integ/Equiv.thy Wed Sep 01 15:03:41 2004 +0200 +++ b/src/HOL/Integ/Equiv.thy Wed Sep 01 15:04:01 2004 +0200 @@ -151,12 +151,19 @@ fixes r and f assumes congruent: "(y,z) \ r ==> f y = f z" +syntax + RESPECTS ::"['a => 'b, ('a * 'a) set] => bool" (infixr "respects" 80) + +translations + "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 @{text UN_equiv_class} *} by auto lemma UN_equiv_class: - "equiv A r ==> congruent r f ==> a \ A + "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+) @@ -165,7 +172,7 @@ done lemma UN_equiv_class_type: - "equiv A r ==> congruent r f ==> X \ A//r ==> + "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 @@ -180,7 +187,7 @@ *} lemma UN_equiv_class_inject: - "equiv A r ==> congruent r f ==> + "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" @@ -203,6 +210,13 @@ assumes congruent2: "(y1,z1) \ r1 ==> (y2,z2) \ r2 ==> f y1 y2 = f z1 z2" +text{*Abbreviation for the common case where the relations are identical*} +syntax + RESPECTS2 ::"['a => 'b, ('a * 'a) set] => bool" (infixr "respects2 " 80) + +translations + "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_def) blast @@ -258,7 +272,7 @@ 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 "congruent2 r r f" + shows "f respects2 r" apply (rule congruent2I [OF equivA equivA]) apply (rule commute [THEN trans]) apply (rule_tac [3] commute [THEN trans, symmetric]) @@ -270,7 +284,7 @@ subsection {* Cardinality results *} -text {* (suggested by Florian Kammüller) *} +text {*Suggested by Florian Kammüller*} lemma finite_quotient: "finite A ==> r \ A \ A ==> finite (A//r)" -- {* recall @{thm equiv_type} *}