new theorem Collect_imp_eq
authorpaulson
Fri, 28 May 2004 11:19:15 +0200
changeset 14812 4b2c006d0534
parent 14811 9144ec118703
child 14813 826edbc317e6
new theorem Collect_imp_eq
src/HOL/Set.thy
--- a/src/HOL/Set.thy	Wed May 26 19:06:09 2004 +0200
+++ b/src/HOL/Set.thy	Fri May 28 11:19:15 2004 +0200
@@ -1108,6 +1108,9 @@
 lemma Collect_disj_eq: "{x. P x | Q x} = {x. P x} \<union> {x. Q x}"
   by blast
 
+lemma Collect_imp_eq: "{x. P x --> Q x} = -{x. P x} \<union> {x. Q x}"
+  by blast
+
 lemma Collect_conj_eq: "{x. P x & Q x} = {x. P x} \<inter> {x. Q x}"
   by blast