| author | huffman | 
| Wed, 17 Aug 2011 15:03:30 -0700 | |
| changeset 44261 | e44f465c00a1 | 
| parent 42871 | 1c0b99f950d9 | 
| child 44326 | 2b088d74beb3 | 
| permissions | -rw-r--r-- | 
| 31807 | 1 | |
| 2 | (* Author: Florian Haftmann, TU Muenchen *) | |
| 3 | ||
| 4 | header {* Relating (finite) sets and lists *}
 | |
| 5 | ||
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 6 | theory More_Set | 
| 37023 | 7 | imports Main More_List | 
| 31807 | 8 | begin | 
| 9 | ||
| 10 | subsection {* Various additional set functions *}
 | |
| 11 | ||
| 12 | definition is_empty :: "'a set \<Rightarrow> bool" where | |
| 13 |   "is_empty A \<longleftrightarrow> A = {}"
 | |
| 14 | ||
| 15 | definition remove :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where | |
| 16 |   "remove x A = A - {x}"
 | |
| 17 | ||
| 42871 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42868diff
changeset | 18 | lemma comp_fun_idem_remove: | 
| 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42868diff
changeset | 19 | "comp_fun_idem remove" | 
| 31807 | 20 | proof - | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 21 |   have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
 | 
| 42871 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42868diff
changeset | 22 | show ?thesis by (simp only: comp_fun_idem_remove rem) | 
| 31807 | 23 | qed | 
| 24 | ||
| 25 | lemma minus_fold_remove: | |
| 26 | assumes "finite A" | |
| 37023 | 27 | shows "B - A = Finite_Set.fold remove B A" | 
| 31807 | 28 | proof - | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 29 |   have rem: "remove = (\<lambda>x A. A - {x})" by (simp add: fun_eq_iff remove_def)
 | 
| 31807 | 30 | show ?thesis by (simp only: rem assms minus_fold_remove) | 
| 31 | qed | |
| 32 | ||
| 33 | definition project :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
 | |
| 34 |   "project P A = {a\<in>A. P a}"
 | |
| 35 | ||
| 36 | ||
| 37 | subsection {* Basic set operations *}
 | |
| 38 | ||
| 39 | lemma is_empty_set: | |
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37024diff
changeset | 40 | "is_empty (set xs) \<longleftrightarrow> List.null xs" | 
| 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37024diff
changeset | 41 | by (simp add: is_empty_def null_def) | 
| 31807 | 42 | |
| 43 | lemma empty_set: | |
| 44 |   "{} = set []"
 | |
| 45 | by simp | |
| 46 | ||
| 32880 | 47 | lemma insert_set_compl: | 
| 34977 | 48 | "insert x (- set xs) = - set (removeAll x xs)" | 
| 49 | by auto | |
| 31807 | 50 | |
| 32880 | 51 | lemma remove_set_compl: | 
| 34977 | 52 | "remove x (- set xs) = - set (List.insert x xs)" | 
| 53 | by (auto simp del: mem_def simp add: remove_def List.insert_def) | |
| 32880 | 54 | |
| 31807 | 55 | lemma image_set: | 
| 31846 | 56 | "image f (set xs) = set (map f xs)" | 
| 31807 | 57 | by simp | 
| 58 | ||
| 59 | lemma project_set: | |
| 60 | "project P (set xs) = set (filter P xs)" | |
| 61 | by (auto simp add: project_def) | |
| 62 | ||
| 63 | ||
| 64 | subsection {* Functorial set operations *}
 | |
| 65 | ||
| 66 | lemma union_set: | |
| 37023 | 67 | "set xs \<union> A = fold Set.insert xs A" | 
| 31807 | 68 | proof - | 
| 42871 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42868diff
changeset | 69 | interpret comp_fun_idem Set.insert | 
| 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42868diff
changeset | 70 | by (fact comp_fun_idem_insert) | 
| 31807 | 71 | show ?thesis by (simp add: union_fold_insert fold_set) | 
| 72 | qed | |
| 73 | ||
| 37023 | 74 | lemma union_set_foldr: | 
| 75 | "set xs \<union> A = foldr Set.insert xs A" | |
| 76 | proof - | |
| 77 | have "\<And>x y :: 'a. insert y \<circ> insert x = insert x \<circ> insert y" | |
| 78 | by (auto intro: ext) | |
| 79 | then show ?thesis by (simp add: union_set foldr_fold) | |
| 80 | qed | |
| 81 | ||
| 31807 | 82 | lemma minus_set: | 
| 37023 | 83 | "A - set xs = fold remove xs A" | 
| 31807 | 84 | proof - | 
| 42871 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42868diff
changeset | 85 | interpret comp_fun_idem remove | 
| 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42868diff
changeset | 86 | by (fact comp_fun_idem_remove) | 
| 31807 | 87 | show ?thesis | 
| 88 | by (simp add: minus_fold_remove [of _ A] fold_set) | |
| 89 | qed | |
| 90 | ||
| 37023 | 91 | lemma minus_set_foldr: | 
| 92 | "A - set xs = foldr remove xs A" | |
| 93 | proof - | |
| 94 | have "\<And>x y :: 'a. remove y \<circ> remove x = remove x \<circ> remove y" | |
| 95 | by (auto simp add: remove_def intro: ext) | |
| 96 | then show ?thesis by (simp add: minus_set foldr_fold) | |
| 97 | qed | |
| 98 | ||
| 31807 | 99 | |
| 100 | subsection {* Derived set operations *}
 | |
| 101 | ||
| 102 | lemma member: | |
| 103 | "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. a = x)" | |
| 104 | by simp | |
| 105 | ||
| 106 | lemma subset_eq: | |
| 107 | "A \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)" | |
| 108 | by (fact subset_eq) | |
| 109 | ||
| 110 | lemma subset: | |
| 111 | "A \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> \<not> B \<subseteq> A" | |
| 112 | by (fact less_le_not_le) | |
| 113 | ||
| 114 | lemma set_eq: | |
| 115 | "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A" | |
| 116 | by (fact eq_iff) | |
| 117 | ||
| 118 | lemma inter: | |
| 119 | "A \<inter> B = project (\<lambda>x. x \<in> A) B" | |
| 120 | by (auto simp add: project_def) | |
| 121 | ||
| 37023 | 122 | |
| 123 | subsection {* Various lemmas *}
 | |
| 124 | ||
| 125 | lemma not_set_compl: | |
| 126 | "Not \<circ> set xs = - set xs" | |
| 42868 | 127 | by (simp add: fun_Compl_def comp_def fun_eq_iff) | 
| 37023 | 128 | |
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 129 | end |