added lemmas
authornipkow
Mon, 02 Sep 2013 17:12:59 +0200
changeset 53364 a4fff0c0599c
parent 53360 7ffc4a746a73
child 53365 643e1151ed7e
added lemmas
src/HOL/Set.thy
--- a/src/HOL/Set.thy	Mon Sep 02 15:13:00 2013 +0200
+++ b/src/HOL/Set.thy	Mon Sep 02 17:12:59 2013 +0200
@@ -883,6 +883,13 @@
 lemma doubleton_eq_iff: "({a,b} = {c,d}) = (a=c & b=d | a=d & b=c)"
   by (blast elim: equalityE)
 
+lemma Un_singleton_iff:
+  "(A \<union> B = {x}) = (A = {} \<and> B = {x} \<or> A = {x} \<and> B = {} \<or> A = {x} \<and> B = {x})"
+by auto
+
+lemma singleton_Un_iff:
+  "({x} = A \<union> B) = (A = {} \<and> B = {x} \<or> A = {x} \<and> B = {} \<or> A = {x} \<and> B = {x})"
+by auto
 
 subsubsection {* Image of a set under a function *}