diff -r de2dc5f5277d -r 48c534b22040 src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Tue Jan 10 17:14:25 2012 +0100 +++ b/src/HOL/Predicate.thy Tue Jan 10 18:09:09 2012 +0100 @@ -73,7 +73,7 @@ subsubsection {* Equality *} -lemma pred_equals_eq: "((\x. x \ R) = (\x. x \ S)) \ (R = S)" +lemma pred_equals_eq [pred_set_conv]: "((\x. x \ R) = (\x. x \ S)) \ (R = S)" by (simp add: set_eq_iff fun_eq_iff) lemma pred_equals_eq2 [pred_set_conv]: "((\x y. (x, y) \ R) = (\x y. (x, y) \ S)) \ (R = S)" @@ -82,7 +82,7 @@ subsubsection {* Order relation *} -lemma pred_subset_eq: "((\x. x \ R) \ (\x. x \ S)) \ (R \ S)" +lemma pred_subset_eq [pred_set_conv]: "((\x. x \ R) \ (\x. x \ S)) \ (R \ S)" by (simp add: subset_iff le_fun_def) lemma pred_subset_eq2 [pred_set_conv]: "((\x y. (x, y) \ R) \ (\x y. (x, y) \ S)) \ (R \ S)" @@ -174,7 +174,7 @@ lemma sup2CI [intro!]: "(\ B x y \ A x y) \ (A \ B) x y" by (auto simp add: sup_fun_def) -lemma sup_Un_eq: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" +lemma sup_Un_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" by (simp add: sup_fun_def) lemma sup_Un_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) \ (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" @@ -234,10 +234,10 @@ lemma SUP2_E [elim!]: "(\x\A. B x) b c \ (\x. x \ A \ B x b c \ R) \ R" by (auto simp add: SUP_apply) -lemma SUP_UN_eq: "(\i. (\x. x \ r i)) = (\x. x \ (\i. r i))" +lemma SUP_UN_eq [pred_set_conv]: "(\i. (\x. x \ r i)) = (\x. x \ (\i. r i))" by (simp add: SUP_apply fun_eq_iff) -lemma SUP_UN_eq2: "(\i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i. r i))" +lemma SUP_UN_eq2 [pred_set_conv]: "(\i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i. r i))" by (simp add: SUP_apply fun_eq_iff) @@ -341,7 +341,7 @@ lemma Powp_Pow_eq [pred_set_conv]: "Powp (\x. x \ A) = (\x. x \ Pow A)" by (auto simp add: Powp_def fun_eq_iff) -lemmas Powp_mono [mono] = Pow_mono [to_pred pred_subset_eq] +lemmas Powp_mono [mono] = Pow_mono [to_pred] subsubsection {* Properties of relations *}