moved various set operations to theory Set (resp. Product_Type)
authorhaftmann
Mon, 26 Dec 2011 18:32:43 +0100
changeset 45986 c9e50153e5ae
parent 45985 2d399a776de2
child 45987 9ba44b49859b
moved various set operations to theory Set (resp. Product_Type)
src/HOL/Library/Cset.thy
src/HOL/Library/More_Set.thy
src/HOL/MicroJava/BV/BVExample.thy
src/HOL/Product_Type.thy
src/HOL/Quotient_Examples/List_Quotient_Cset.thy
src/HOL/Quotient_Examples/Quotient_Cset.thy
src/HOL/Set.thy
--- a/src/HOL/Library/Cset.thy	Mon Dec 26 17:40:43 2011 +0100
+++ b/src/HOL/Library/Cset.thy	Mon Dec 26 18:32:43 2011 +0100
@@ -103,13 +103,13 @@
 hide_const (open) UNIV
 
 definition is_empty :: "'a Cset.set \<Rightarrow> bool" where
-  [simp]: "is_empty A \<longleftrightarrow> More_Set.is_empty (set_of A)"
+  [simp]: "is_empty A \<longleftrightarrow> Set.is_empty (set_of A)"
 
 definition insert :: "'a \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
   [simp]: "insert x A = Set (Set.insert x (set_of A))"
 
 definition remove :: "'a \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
-  [simp]: "remove x A = Set (More_Set.remove x (set_of A))"
+  [simp]: "remove x A = Set (Set.remove x (set_of A))"
 
 definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a Cset.set \<Rightarrow> 'b Cset.set" where
   [simp]: "map f A = Set (image f (set_of A))"
@@ -118,7 +118,7 @@
   by (simp_all add: fun_eq_iff image_compose)
 
 definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
-  [simp]: "filter P A = Set (More_Set.project P (set_of A))"
+  [simp]: "filter P A = Set (Set.project P (set_of A))"
 
 definition forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> bool" where
   [simp]: "forall P A \<longleftrightarrow> Ball (set_of A) P"
@@ -182,17 +182,17 @@
 
 lemma is_empty_simp [simp]:
   "is_empty A \<longleftrightarrow> set_of A = {}"
-  by (simp add: More_Set.is_empty_def)
+  by (simp add: Set.is_empty_def)
 declare is_empty_def [simp del]
 
 lemma remove_simp [simp]:
   "remove x A = Set (set_of A - {x})"
-  by (simp add: More_Set.remove_def)
+  by (simp add: Set.remove_def)
 declare remove_def [simp del]
 
 lemma filter_simp [simp]:
   "filter P A = Set {x \<in> set_of A. P x}"
-  by (simp add: More_Set.project_def)
+  by (simp add: Set.project_def)
 declare filter_def [simp del]
 
 lemma set_of_set [simp]:
@@ -297,18 +297,18 @@
 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)
-  have *: "\<And>x::'a. Cset.remove = (\<lambda>x. Set \<circ> More_Set.remove x \<circ> set_of)"
-    by (simp add: fun_eq_iff More_Set.remove_def)
-  have "set_of \<circ> fold (\<lambda>x. Set \<circ> More_Set.remove x \<circ> set_of) xs =
-    fold More_Set.remove xs \<circ> set_of"
+  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 =
+    fold Set.remove xs \<circ> set_of"
     by (rule fold_commute) (simp add: fun_eq_iff)
-  then have "fold More_Set.remove xs (set_of A) = 
-    set_of (fold (\<lambda>x. Set \<circ> More_Set.remove x \<circ> set_of) xs A)"
+  then have "fold Set.remove xs (set_of A) = 
+    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 *)
   moreover have "\<And>x y :: 'a. Cset.remove y \<circ> Cset.remove x = Cset.remove x \<circ> Cset.remove y"
-    by (auto simp add: More_Set.remove_def *)
+    by (auto simp add: Set.remove_def *)
   ultimately show "inf A (Cset.coset xs) = foldr Cset.remove xs A"
     by (simp add: foldr_fold)
 qed
--- a/src/HOL/Library/More_Set.thy	Mon Dec 26 17:40:43 2011 +0100
+++ b/src/HOL/Library/More_Set.thy	Mon Dec 26 18:32:43 2011 +0100
@@ -7,57 +7,31 @@
 imports Main More_List
 begin
 
-subsection {* Various additional set functions *}
-
-definition is_empty :: "'a set \<Rightarrow> bool" where
-  "is_empty A \<longleftrightarrow> A = {}"
-
-hide_const (open) is_empty
-
-definition remove :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where
-  "remove x A = A - {x}"
-
-hide_const (open) remove
-
 lemma comp_fun_idem_remove:
-  "comp_fun_idem More_Set.remove"
+  "comp_fun_idem Set.remove"
 proof -
-  have rem: "More_Set.remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
+  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 More_Set.remove B A"
+  shows "B - A = Finite_Set.fold Set.remove B A"
 proof -
-  have rem: "More_Set.remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
+  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
 
-definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
-  "project P A = {a \<in> A. P a}"
-
-hide_const (open) project
-
 lemma bounded_Collect_code [code_unfold_post]:
-  "{x \<in> A. P x} = More_Set.project P A"
+  "{x \<in> A. P x} = Set.project P A"
   by (simp add: project_def)
 
-definition product :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
-  "product A B = Sigma A (\<lambda>_. B)"
-
-hide_const (open) product
-
-lemma [code_unfold_post]:
-  "Sigma A (\<lambda>_. B) = More_Set.product A B"
-  by (simp add: product_def)
-
 
 subsection {* Basic set operations *}
 
 lemma is_empty_set:
-  "More_Set.is_empty (set xs) \<longleftrightarrow> List.null xs"
-  by (simp add: is_empty_def null_def)
+  "Set.is_empty (set xs) \<longleftrightarrow> List.null xs"
+  by (simp add: Set.is_empty_def null_def)
 
 lemma empty_set:
   "{} = set []"
@@ -68,7 +42,7 @@
   by auto
 
 lemma remove_set_compl:
-  "More_Set.remove x (- set xs) = - set (List.insert x xs)"
+  "Set.remove x (- set xs) = - set (List.insert x xs)"
   by (auto simp add: remove_def List.insert_def)
 
 lemma image_set:
@@ -76,7 +50,7 @@
   by simp
 
 lemma project_set:
-  "More_Set.project P (set xs) = set (filter P xs)"
+  "Set.project P (set xs) = set (filter P xs)"
   by (auto simp add: project_def)
 
 
@@ -99,18 +73,18 @@
 qed
 
 lemma minus_set:
-  "A - set xs = fold More_Set.remove xs A"
+  "A - set xs = fold Set.remove xs A"
 proof -
-  interpret comp_fun_idem More_Set.remove
+  interpret comp_fun_idem Set.remove
     by (fact comp_fun_idem_remove)
   show ?thesis
     by (simp add: minus_fold_remove [of _ A] fold_set)
 qed
 
 lemma minus_set_foldr:
-  "A - set xs = foldr More_Set.remove xs A"
+  "A - set xs = foldr Set.remove xs A"
 proof -
-  have "\<And>x y :: 'a. More_Set.remove y \<circ> More_Set.remove x = More_Set.remove x \<circ> More_Set.remove y"
+  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)
 qed
@@ -135,15 +109,15 @@
   by (fact eq_iff)
 
 lemma inter:
-  "A \<inter> B = More_Set.project (\<lambda>x. x \<in> A) B"
+  "A \<inter> B = Set.project (\<lambda>x. x \<in> A) B"
   by (auto simp add: project_def)
 
 
 subsection {* Theorems on relations *}
 
 lemma product_code:
-  "More_Set.product (set xs) (set ys) = set [(x, y). x \<leftarrow> xs, y \<leftarrow> ys]"
-  by (auto simp add: product_def)
+  "Product_Type.product (set xs) (set ys) = set [(x, y). x \<leftarrow> xs, y \<leftarrow> ys]"
+  by (auto simp add: Product_Type.product_def)
 
 lemma Id_on_set:
   "Id_on (set xs) = set [(x, x). x \<leftarrow> xs]"
@@ -177,8 +151,8 @@
   by (simp_all add: member_def)
 
 lemma [code_unfold]:
-  "A = {} \<longleftrightarrow> More_Set.is_empty A"
-  by (simp add: is_empty_def)
+  "A = {} \<longleftrightarrow> Set.is_empty A"
+  by (simp add: Set.is_empty_def)
 
 declare empty_set [code]
 
@@ -194,8 +168,8 @@
   by simp_all
 
 lemma remove_code [code]:
-  "More_Set.remove x (set xs) = set (removeAll x xs)"
-  "More_Set.remove x (coset xs) = coset (List.insert x xs)"
+  "Set.remove x (set xs) = set (removeAll x xs)"
+  "Set.remove x (coset xs) = coset (List.insert x xs)"
   by (simp_all add: remove_def Compl_insert)
 
 declare image_set [code]
@@ -221,17 +195,6 @@
 
 subsection {* Derived operations *}
 
-instantiation set :: (equal) equal
-begin
-
-definition
-  "HOL.equal A B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
-
-instance proof
-qed (auto simp add: equal_set_def)
-
-end
-
 declare subset_eq [code]
 
 declare subset [code]
@@ -241,11 +204,11 @@
 
 lemma inter_code [code]:
   "A \<inter> set xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
-  "A \<inter> coset xs = foldr More_Set.remove xs A"
+  "A \<inter> coset xs = foldr Set.remove xs A"
   by (simp add: inter project_def) (simp add: Diff_eq [symmetric] minus_set_foldr)
 
 lemma subtract_code [code]:
-  "A - set xs = foldr More_Set.remove xs A"
+  "A - set xs = foldr Set.remove xs A"
   "A - coset xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
   by (auto simp add: minus_set_foldr)
 
--- a/src/HOL/MicroJava/BV/BVExample.thy	Mon Dec 26 17:40:43 2011 +0100
+++ b/src/HOL/MicroJava/BV/BVExample.thy	Mon Dec 26 18:32:43 2011 +0100
@@ -451,11 +451,11 @@
 qed
 
 lemma [code]:
-  "iter f step ss w = while (\<lambda>(ss, w). \<not> More_Set.is_empty w)
+  "iter f step ss w = while (\<lambda>(ss, w). \<not> Set.is_empty w)
     (\<lambda>(ss, w).
         let p = some_elem w in propa f (step p (ss ! p)) ss (w - {p}))
     (ss, w)"
-  unfolding iter_def More_Set.is_empty_def some_elem_def ..
+  unfolding iter_def Set.is_empty_def some_elem_def ..
 
 lemma JVM_sup_unfold [code]:
  "JVMType.sup S m n = lift2 (Opt.sup
@@ -469,14 +469,6 @@
 lemmas [code] = JType.sup_def [unfolded exec_lub_def] JVM_le_unfold
 lemmas [code] = lesub_def plussub_def
 
-setup {*
-  Code.add_signature_cmd ("More_Set.is_empty", "'a set \<Rightarrow> bool")
-  #> Code.add_signature_cmd ("propa", "('s \<Rightarrow> 's \<Rightarrow> 's) \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list * nat set") 
-  #> Code.add_signature_cmd 
-     ("iter", "('s \<Rightarrow> 's \<Rightarrow> 's) \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> (nat \<times> 's) list) \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list * nat set")
-  #> Code.add_signature_cmd ("unstables", "('s \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> (nat \<times> 's) list) \<Rightarrow> 's list \<Rightarrow> nat set") 
-*}
-
 lemmas [code] =
   JType.sup_def [unfolded exec_lub_def]
   wf_class_code
--- a/src/HOL/Product_Type.thy	Mon Dec 26 17:40:43 2011 +0100
+++ b/src/HOL/Product_Type.thy	Mon Dec 26 18:32:43 2011 +0100
@@ -877,11 +877,11 @@
   Disjoint union of a family of sets -- Sigma.
 *}
 
-definition Sigma :: "['a set, 'a => 'b set] => ('a \<times> 'b) set" where
+definition Sigma :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> ('a \<times> 'b) set" where
   Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}"
 
 abbreviation
-  Times :: "['a set, 'b set] => ('a * 'b) set"
+  Times :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set"
     (infixr "<*>" 80) where
   "A <*> B == Sigma A (%_. B)"
 
@@ -893,6 +893,15 @@
 
 hide_const (open) Times
 
+definition product :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
+  "product A B = Sigma A (\<lambda>_. B)"
+
+hide_const (open) product
+
+lemma [code_unfold_post]:
+  "Sigma A (\<lambda>_. B) = Product_Type.product A B"
+  by (simp add: product_def)
+
 syntax
   "_Sigma" :: "[pttrn, 'a set, 'b set] => ('a * 'b) set"  ("(3SIGMA _:_./ _)" [0, 0, 10] 10)
 translations
--- a/src/HOL/Quotient_Examples/List_Quotient_Cset.thy	Mon Dec 26 17:40:43 2011 +0100
+++ b/src/HOL/Quotient_Examples/List_Quotient_Cset.thy	Mon Dec 26 18:32:43 2011 +0100
@@ -75,7 +75,7 @@
   "Quotient_Cset.remove x (coset xs) = coset (List.insert x xs)"
 unfolding coset_def
 apply descending
-apply (simp add: More_Set.remove_def)
+apply (simp add: Set.remove_def)
 apply descending
 by (simp add: remove_set_compl)
 
--- a/src/HOL/Quotient_Examples/Quotient_Cset.thy	Mon Dec 26 17:40:43 2011 +0100
+++ b/src/HOL/Quotient_Examples/Quotient_Cset.thy	Mon Dec 26 18:32:43 2011 +0100
@@ -24,11 +24,11 @@
 lemma [quot_respect]:
   "(op = ===> set_eq ===> op =) (op \<in>) (op \<in>)"
   "(op = ===> set_eq) Collect Collect"
-  "(set_eq ===> op =) More_Set.is_empty More_Set.is_empty"
+  "(set_eq ===> op =) Set.is_empty Set.is_empty"
   "(op = ===> set_eq ===> set_eq) Set.insert Set.insert"
-  "(op = ===> set_eq ===> set_eq) More_Set.remove More_Set.remove"
+  "(op = ===> set_eq ===> set_eq) Set.remove Set.remove"
   "(op = ===> set_eq ===> set_eq) image image"
-  "(op = ===> set_eq ===> set_eq) More_Set.project More_Set.project"
+  "(op = ===> set_eq ===> set_eq) Set.project Set.project"
   "(set_eq ===> op =) Ball Ball"
   "(set_eq ===> op =) Bex Bex"
   "(set_eq ===> op =) Finite_Set.card Finite_Set.card"
@@ -51,15 +51,15 @@
 quotient_definition "Set :: ('a => bool) => 'a Quotient_Cset.set"
 is Collect
 quotient_definition is_empty where "is_empty :: 'a Quotient_Cset.set \<Rightarrow> bool"
-is More_Set.is_empty
+is Set.is_empty
 quotient_definition insert where "insert :: 'a \<Rightarrow> 'a Quotient_Cset.set \<Rightarrow> 'a Quotient_Cset.set"
 is Set.insert
 quotient_definition remove where "remove :: 'a \<Rightarrow> 'a Quotient_Cset.set \<Rightarrow> 'a Quotient_Cset.set"
-is More_Set.remove
+is Set.remove
 quotient_definition map where "map :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a Quotient_Cset.set \<Rightarrow> 'b Quotient_Cset.set"
 is image
 quotient_definition filter where "filter :: ('a \<Rightarrow> bool) \<Rightarrow> 'a Quotient_Cset.set \<Rightarrow> 'a Quotient_Cset.set"
-is More_Set.project
+is Set.project
 quotient_definition "forall :: 'a Quotient_Cset.set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 is Ball
 quotient_definition "exists :: 'a Quotient_Cset.set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
--- a/src/HOL/Set.thy	Mon Dec 26 17:40:43 2011 +0100
+++ b/src/HOL/Set.thy	Mon Dec 26 18:32:43 2011 +0100
@@ -1791,6 +1791,34 @@
 hide_const (open) bind
 
 
+subsubsection {* Operations for execution *}
+
+definition is_empty :: "'a set \<Rightarrow> bool" where
+  "is_empty A \<longleftrightarrow> A = {}"
+
+hide_const (open) is_empty
+
+definition remove :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+  "remove x A = A - {x}"
+
+hide_const (open) remove
+
+definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+  "project P A = {a \<in> A. P a}"
+
+hide_const (open) project
+
+instantiation set :: (equal) equal
+begin
+
+definition
+  "HOL.equal A B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
+
+instance proof
+qed (auto simp add: equal_set_def)
+
+end
+
 text {* Misc *}
 
 hide_const (open) member not_member