diff -r 10a6ba364ce1 -r 36d596cb63a2 Quot/Quotient.thy --- a/Quot/Quotient.thy Fri Feb 12 12:06:09 2010 +0100 +++ b/Quot/Quotient.thy Fri Feb 12 15:06:20 2010 +0100 @@ -12,22 +12,23 @@ ("quotient_tacs.ML") begin + text {* Basic definition for equivalence relations that are represented by predicates. *} definition - "equivp E \ (\x y. E x y = (E x = E y))" + "equivp E \ \x y. E x y = (E x = E y)" definition - "reflp E \ (\x. E x x)" + "reflp E \ \x. E x x" definition - "symp E \ (\x y. E x y \ E y x)" + "symp E \ \x y. E x y \ E y x" definition - "transp E \ (\x y z. E x y \ E y z \ E x z)" + "transp E \ \x y z. E x y \ E y z \ E x z" lemma equivp_reflp_symp_transp: shows "equivp E = (reflp E \ symp E \ transp E)" @@ -51,19 +52,15 @@ shows "equivp R" using assms by (simp add: equivp_reflp_symp_transp) -lemma eq_imp_rel: - shows "equivp R \ a = b \ R a b" - by (simp add: equivp_reflp) - lemma identity_equivp: shows "equivp (op =)" unfolding equivp_def by auto +text {* Partial equivalences: not yet used anywhere *} -text {* Partial equivalences: not yet used anywhere *} definition - "part_equivp E \ ((\x. E x x) \ (\x y. E x y = (E x x \ E y y \ (E x = E y))))" + "part_equivp E \ (\x. E x x) \ (\x y. E x y = (E x x \ E y y \ (E x = E y)))" lemma equivp_implies_part_equivp: assumes a: "equivp E" @@ -88,7 +85,7 @@ definition Respects where - "Respects R x \ (R x x)" + "Respects R x \ R x x" lemma in_respects: shows "(x \ Respects R) = R x x" @@ -130,7 +127,7 @@ section {* Quotient Predicate *} definition - "Quotient E Abs Rep \ + "Quotient E Abs Rep \ (\a. Abs (Rep a) = a) \ (\a. E (Rep a) (Rep a)) \ (\r s. E r s = (E r r \ E s s \ (Abs r = Abs s)))" @@ -354,6 +351,7 @@ apply(simp add: reflp_def) done +(* Next four lemmas are unused *) lemma all_reg: assumes a: "!x :: 'a. (P x --> Q x)" and b: "All P" @@ -378,6 +376,7 @@ shows "Bex R Q" using a b by (metis COMBC_def Collect_def Collect_mem_eq) + lemma ball_all_comm: assumes "\y. (\x\P. A x y) \ (\x. B x y)" shows "(\x\P. \y. A x y) \ (\x. \y. B x y)"