incorporated various theorems from theory More_Set into corpus
authorhaftmann
Fri, 06 Jan 2012 21:48:45 +0100
changeset 46146 6baea4fca6bd
parent 46145 0ec0af1c651d
child 46147 2c4d8de91c4c
incorporated various theorems from theory More_Set into corpus
src/HOL/Finite_Set.thy
src/HOL/Library/Cset.thy
src/HOL/More_Set.thy
src/HOL/Set.thy
--- a/src/HOL/Finite_Set.thy	Fri Jan 06 21:48:45 2012 +0100
+++ b/src/HOL/Finite_Set.thy	Fri Jan 06 21:48:45 2012 +0100
@@ -718,7 +718,7 @@
 qed auto
 
 lemma comp_fun_idem_remove:
-  "comp_fun_idem (\<lambda>x A. A - {x})"
+  "comp_fun_idem Set.remove"
 proof
 qed auto
 
@@ -742,10 +742,11 @@
 
 lemma minus_fold_remove:
   assumes "finite A"
-  shows "B - A = fold (\<lambda>x A. A - {x}) B A"
+  shows "B - A = fold Set.remove B A"
 proof -
-  interpret comp_fun_idem "\<lambda>x A. A - {x}" by (fact comp_fun_idem_remove)
-  from `finite A` show ?thesis by (induct A arbitrary: B) auto
+  interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove)
+  from `finite A` have "fold Set.remove B A = B - A" by (induct A arbitrary: B) auto
+  then show ?thesis ..
 qed
 
 context complete_lattice
@@ -779,7 +780,7 @@
   shows "Sup A = fold sup bot A"
   using assms sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2)
 
-lemma inf_INFI_fold_inf:
+lemma inf_INF_fold_inf:
   assumes "finite A"
   shows "inf B (INFI A f) = fold (inf \<circ> f) B A" (is "?inf = ?fold") 
 proof (rule sym)
@@ -790,7 +791,7 @@
       (simp_all add: INF_def inf_left_commute)
 qed
 
-lemma sup_SUPR_fold_sup:
+lemma sup_SUP_fold_sup:
   assumes "finite A"
   shows "sup B (SUPR A f) = fold (sup \<circ> f) B A" (is "?sup = ?fold") 
 proof (rule sym)
@@ -801,15 +802,15 @@
       (simp_all add: SUP_def sup_left_commute)
 qed
 
-lemma INFI_fold_inf:
+lemma INF_fold_inf:
   assumes "finite A"
   shows "INFI A f = fold (inf \<circ> f) top A"
-  using assms inf_INFI_fold_inf [of A top] by simp
+  using assms inf_INF_fold_inf [of A top] by simp
 
-lemma SUPR_fold_sup:
+lemma SUP_fold_sup:
   assumes "finite A"
   shows "SUPR A f = fold (sup \<circ> f) bot A"
-  using assms sup_SUPR_fold_sup [of A bot] by simp
+  using assms sup_SUP_fold_sup [of A bot] by simp
 
 end
 
@@ -2066,10 +2067,10 @@
   assumes fin: "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
   shows "finite (UNIV :: 'b set)"
 proof -
-  from fin have "finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary))"
-    by(rule finite_imageI)
-  moreover have "UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary)"
-    by(rule UNIV_eq_I) auto
+  from fin have "\<And>arbitrary. finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary))"
+    by (rule finite_imageI)
+  moreover have "\<And>arbitrary. UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary)"
+    by (rule UNIV_eq_I) auto
   ultimately show "finite (UNIV :: 'b set)" by simp
 qed
 
--- a/src/HOL/Library/Cset.thy	Fri Jan 06 21:48:45 2012 +0100
+++ b/src/HOL/Library/Cset.thy	Fri Jan 06 21:48:45 2012 +0100
@@ -276,7 +276,7 @@
   fix xs :: "'a list"
   show "member (Cset.set xs) = member (fold insert xs Cset.empty)"
     by (simp add: fold_commute_apply [symmetric, where ?h = Set and ?g = Set.insert]
-      fun_eq_iff Cset.set_def union_set [symmetric])
+      fun_eq_iff Cset.set_def union_set_fold [symmetric])
 qed
 
 lemma single_code [code]:
@@ -296,7 +296,7 @@
   "inf A (Cset.coset xs) = foldr Cset.remove xs A"
 proof -
   show "inf A (Cset.set xs) = Cset.set (List.filter (member A) xs)"
-    by (simp add: inter project_def Cset.set_def member_def)
+    by (simp add: project_def Cset.set_def member_def) auto
   have *: "\<And>x::'a. Cset.remove = (\<lambda>x. Set \<circ> Set.remove x \<circ> set_of)"
     by (simp add: fun_eq_iff Set.remove_def)
   have "set_of \<circ> fold (\<lambda>x. Set \<circ> Set.remove x \<circ> set_of) xs =
@@ -306,7 +306,7 @@
     set_of (fold (\<lambda>x. Set \<circ> Set.remove x \<circ> set_of) xs A)"
     by (simp add: fun_eq_iff)
   then have "inf A (Cset.coset xs) = fold Cset.remove xs A"
-    by (simp add: Diff_eq [symmetric] minus_set *)
+    by (simp add: Diff_eq [symmetric] minus_set_fold *)
   moreover have "\<And>x y :: 'a. Cset.remove y \<circ> Cset.remove x = Cset.remove x \<circ> Cset.remove y"
     by (auto simp add: Set.remove_def *)
   ultimately show "inf A (Cset.coset xs) = foldr Cset.remove xs A"
@@ -326,7 +326,7 @@
     set_of (fold (\<lambda>x. Set \<circ> Set.insert x \<circ> set_of) xs A)"
     by (simp add: fun_eq_iff)
   then have "sup (Cset.set xs) A = fold Cset.insert xs A"
-    by (simp add: union_set *)
+    by (simp add: union_set_fold *)
   moreover have "\<And>x y :: 'a. Cset.insert y \<circ> Cset.insert x = Cset.insert x \<circ> Cset.insert y"
     by (auto simp add: *)
   ultimately show "sup (Cset.set xs) A = foldr Cset.insert xs A"
--- a/src/HOL/More_Set.thy	Fri Jan 06 21:48:45 2012 +0100
+++ b/src/HOL/More_Set.thy	Fri Jan 06 21:48:45 2012 +0100
@@ -7,26 +7,6 @@
 imports List
 begin
 
-lemma comp_fun_idem_remove:
-  "comp_fun_idem Set.remove"
-proof -
-  have rem: "Set.remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
-  show ?thesis by (simp only: comp_fun_idem_remove rem)
-qed
-
-lemma minus_fold_remove:
-  assumes "finite A"
-  shows "B - A = Finite_Set.fold Set.remove B A"
-proof -
-  have rem: "Set.remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
-  show ?thesis by (simp only: rem assms minus_fold_remove)
-qed
-
-lemma bounded_Collect_code: (* FIXME delete candidate *)
-  "{x \<in> A. P x} = Set.project P A"
-  by (simp add: project_def)
-
-
 subsection {* Basic set operations *}
 
 lemma is_empty_set [code]:
@@ -37,26 +17,10 @@
   "{} = set []"
   by simp
 
-lemma insert_set_compl:
-  "insert x (- set xs) = - set (removeAll x xs)"
-  by auto
-
-lemma remove_set_compl:
-  "Set.remove x (- set xs) = - set (List.insert x xs)"
-  by (auto simp add: remove_def List.insert_def)
-
-lemma image_set:
-  "image f (set xs) = set (map f xs)"
-  by simp
-
-lemma project_set:
-  "Set.project P (set xs) = set (filter P xs)"
-  by (auto simp add: project_def)
-
 
 subsection {* Functorial set operations *}
 
-lemma union_set:
+lemma union_set_fold:
   "set xs \<union> A = fold Set.insert xs A"
 proof -
   interpret comp_fun_idem Set.insert
@@ -69,10 +33,10 @@
 proof -
   have "\<And>x y :: 'a. insert y \<circ> insert x = insert x \<circ> insert y"
     by auto
-  then show ?thesis by (simp add: union_set foldr_fold)
+  then show ?thesis by (simp add: union_set_fold foldr_fold)
 qed
 
-lemma minus_set:
+lemma minus_set_fold:
   "A - set xs = fold Set.remove xs A"
 proof -
   interpret comp_fun_idem Set.remove
@@ -86,56 +50,29 @@
 proof -
   have "\<And>x y :: 'a. Set.remove y \<circ> Set.remove x = Set.remove x \<circ> Set.remove y"
     by (auto simp add: remove_def)
-  then show ?thesis by (simp add: minus_set foldr_fold)
+  then show ?thesis by (simp add: minus_set_fold foldr_fold)
 qed
 
 
-subsection {* Derived set operations *}
-
-lemma member [code]:
-  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
-  by simp
-
-lemma subset [code]:
-  "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
-  by (fact less_le_not_le)
-
-lemma set_eq [code]:
-  "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
-  by (fact eq_iff)
-
-lemma inter [code]:
-  "A \<inter> B = Set.project (\<lambda>x. x \<in> A) B"
-  by (auto simp add: project_def)
-
-
-subsection {* Code generator setup *}
-
-definition coset :: "'a list \<Rightarrow> 'a set" where
-  [simp]: "coset xs = - set xs"
-
-code_datatype set coset
-
-
 subsection {* Basic operations *}
 
 lemma [code]:
   "x \<in> set xs \<longleftrightarrow> List.member xs x"
-  "x \<in> coset xs \<longleftrightarrow> \<not> List.member xs x"
+  "x \<in> List.coset xs \<longleftrightarrow> \<not> List.member xs x"
   by (simp_all add: member_def)
 
 lemma UNIV_coset [code]:
-  "UNIV = coset []"
+  "UNIV = List.coset []"
   by simp
 
 lemma insert_code [code]:
   "insert x (set xs) = set (List.insert x xs)"
-  "insert x (coset xs) = coset (removeAll x xs)"
+  "insert x (List.coset xs) = List.coset (removeAll x xs)"
   by simp_all
 
 lemma remove_code [code]:
   "Set.remove x (set xs) = set (removeAll x xs)"
-  "Set.remove x (coset xs) = coset (List.insert x xs)"
+  "Set.remove x (List.coset xs) = List.coset (List.insert x xs)"
   by (simp_all add: remove_def Compl_insert)
 
 lemma Ball_set [code]:
@@ -159,17 +96,17 @@
 
 lemma inter_code [code]:
   "A \<inter> set xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
-  "A \<inter> coset xs = foldr Set.remove xs A"
-  by (simp add: inter project_def) (simp add: Diff_eq [symmetric] minus_set_foldr)
+  "A \<inter> List.coset xs = foldr Set.remove xs A"
+  by (simp add: inter_project project_def) (simp add: Diff_eq [symmetric] minus_set_foldr)
 
 lemma subtract_code [code]:
   "A - set xs = foldr Set.remove xs A"
-  "A - coset xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
+  "A - List.coset xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
   by (auto simp add: minus_set_foldr)
 
 lemma union_code [code]:
   "set xs \<union> A = foldr insert xs A"
-  "coset xs \<union> A = coset (List.filter (\<lambda>x. x \<notin> A) xs)"
+  "List.coset xs \<union> A = List.coset (List.filter (\<lambda>x. x \<notin> A) xs)"
   by (auto simp add: union_set_foldr)
 
 definition Inf :: "'a::complete_lattice set \<Rightarrow> 'a" where
@@ -184,12 +121,12 @@
 
 lemma Inf_code [code]:
   "More_Set.Inf (set xs) = foldr inf xs top"
-  "More_Set.Inf (coset []) = bot"
+  "More_Set.Inf (List.coset []) = bot"
   by (simp_all add: Inf_set_foldr)
 
 lemma Sup_sup [code]:
   "More_Set.Sup (set xs) = foldr sup xs bot"
-  "More_Set.Sup (coset []) = top"
+  "More_Set.Sup (List.coset []) = top"
   by (simp_all add: Sup_set_foldr)
 
 (* FIXME: better implement conversion by bisection *)
@@ -226,7 +163,8 @@
   "Id_on (set xs) = set [(x, x). x \<leftarrow> xs]"
   by (auto simp add: Id_on_def)
 
-lemma trancl_set_ntrancl [code]: "trancl (set xs) = ntrancl (card (set xs) - 1) (set xs)"
+lemma trancl_set_ntrancl [code]:
+  "trancl (set xs) = ntrancl (card (set xs) - 1) (set xs)"
   by (simp add: finite_trancl_ntranl)
 
 lemma set_rel_comp [code]:
--- a/src/HOL/Set.thy	Fri Jan 06 21:48:45 2012 +0100
+++ b/src/HOL/Set.thy	Fri Jan 06 21:48:45 2012 +0100
@@ -431,6 +431,10 @@
 lemma bex_triv_one_point2 [simp]: "(EX x:A. a = x) = (a:A)"
   by blast
 
+lemma member_exists [code]:
+  "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)"
+  by (rule sym) (fact bex_triv_one_point2)
+
 lemma bex_one_point1 [simp]: "(EX x:A. x = a & P x) = (a:A & P a)"
   by blast
 
@@ -522,6 +526,10 @@
 lemma set_mp: "A \<subseteq> B ==> x:A ==> x:B"
   by (rule subsetD)
 
+lemma subset_not_subset_eq [code]:
+  "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A"
+  by (fact less_le_not_le)
+
 lemma eq_mem_trans: "a=b ==> b \<in> A ==> a \<in> A"
   by simp
 
@@ -1829,6 +1837,10 @@
   "x \<in> Set.project P A \<longleftrightarrow> x \<in> A \<and> P x"
   by (simp add: project_def)
 
+lemma inter_project [code]:
+  "A \<inter> B = Set.project (\<lambda>x. x \<in> A) B"
+  by (auto simp add: project_def)
+
 instantiation set :: (equal) equal
 begin