| author | nipkow | 
| Wed, 24 Aug 2016 11:02:23 +0200 | |
| changeset 63720 | bcf2123d059a | 
| parent 63148 | 6a767355d1a9 | 
| child 63928 | d81fb5b46a5c | 
| permissions | -rw-r--r-- | 
| 60727 | 1 | (* Title: HOL/Library/Disjoint_Sets.thy | 
| 2 | Author: Johannes Hölzl, TU München | |
| 3 | *) | |
| 4 | ||
| 63098 | 5 | section \<open>Partitions and Disjoint Sets\<close> | 
| 60727 | 6 | |
| 7 | theory Disjoint_Sets | |
| 8 | imports Main | |
| 9 | begin | |
| 10 | ||
| 11 | lemma range_subsetD: "range f \<subseteq> B \<Longrightarrow> f i \<in> B" | |
| 12 | by blast | |
| 13 | ||
| 14 | lemma Int_Diff_disjoint: "A \<inter> B \<inter> (A - B) = {}"
 | |
| 15 | by blast | |
| 16 | ||
| 17 | lemma Int_Diff_Un: "A \<inter> B \<union> (A - B) = A" | |
| 18 | by blast | |
| 19 | ||
| 20 | lemma mono_Un: "mono A \<Longrightarrow> (\<Union>i\<le>n. A i) = A n" | |
| 21 | unfolding mono_def by auto | |
| 22 | ||
| 63098 | 23 | lemma disjnt_equiv_class: "equiv A r \<Longrightarrow> disjnt (r``{a}) (r``{b}) \<longleftrightarrow> (a, b) \<notin> r"
 | 
| 24 | by (auto dest: equiv_class_self simp: equiv_class_eq_iff disjnt_def) | |
| 25 | ||
| 60727 | 26 | subsection \<open>Set of Disjoint Sets\<close> | 
| 27 | ||
| 62843 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 28 | abbreviation disjoint :: "'a set set \<Rightarrow> bool" where "disjoint \<equiv> pairwise disjnt" | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 29 | |
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 30 | lemma disjoint_def: "disjoint A \<longleftrightarrow> (\<forall>a\<in>A. \<forall>b\<in>A. a \<noteq> b \<longrightarrow> a \<inter> b = {})"
 | 
| 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 31 | unfolding pairwise_def disjnt_def by auto | 
| 60727 | 32 | |
| 33 | lemma disjointI: | |
| 34 |   "(\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> a \<inter> b = {}) \<Longrightarrow> disjoint A"
 | |
| 35 | unfolding disjoint_def by auto | |
| 36 | ||
| 37 | lemma disjointD: | |
| 38 |   "disjoint A \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> a \<inter> b = {}"
 | |
| 39 | unfolding disjoint_def by auto | |
| 40 | ||
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 41 | lemma disjoint_image: "inj_on f (\<Union>A) \<Longrightarrow> disjoint A \<Longrightarrow> disjoint (op ` f ` A)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 42 | unfolding inj_on_def disjoint_def by blast | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 43 | |
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 44 | lemma assumes "disjoint (A \<union> B)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 45 | shows disjoint_unionD1: "disjoint A" and disjoint_unionD2: "disjoint B" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 46 | using assms by (simp_all add: disjoint_def) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 47 | |
| 60727 | 48 | lemma disjoint_INT: | 
| 49 | assumes *: "\<And>i. i \<in> I \<Longrightarrow> disjoint (F i)" | |
| 50 |   shows "disjoint {\<Inter>i\<in>I. X i | X. \<forall>i\<in>I. X i \<in> F i}"
 | |
| 51 | proof (safe intro!: disjointI del: equalityI) | |
| 63098 | 52 | fix A B :: "'a \<Rightarrow> 'b set" assume "(\<Inter>i\<in>I. A i) \<noteq> (\<Inter>i\<in>I. B i)" | 
| 60727 | 53 | then obtain i where "A i \<noteq> B i" "i \<in> I" | 
| 54 | by auto | |
| 55 | moreover assume "\<forall>i\<in>I. A i \<in> F i" "\<forall>i\<in>I. B i \<in> F i" | |
| 56 |   ultimately show "(\<Inter>i\<in>I. A i) \<inter> (\<Inter>i\<in>I. B i) = {}"
 | |
| 57 | using *[OF \<open>i\<in>I\<close>, THEN disjointD, of "A i" "B i"] | |
| 58 | by (auto simp: INT_Int_distrib[symmetric]) | |
| 59 | qed | |
| 60 | ||
| 61 | subsubsection "Family of Disjoint Sets" | |
| 62 | ||
| 63 | definition disjoint_family_on :: "('i \<Rightarrow> 'a set) \<Rightarrow> 'i set \<Rightarrow> bool" where
 | |
| 64 |   "disjoint_family_on A S \<longleftrightarrow> (\<forall>m\<in>S. \<forall>n\<in>S. m \<noteq> n \<longrightarrow> A m \<inter> A n = {})"
 | |
| 65 | ||
| 66 | abbreviation "disjoint_family A \<equiv> disjoint_family_on A UNIV" | |
| 67 | ||
| 68 | lemma disjoint_family_onD: | |
| 69 |   "disjoint_family_on A I \<Longrightarrow> i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> i \<noteq> j \<Longrightarrow> A i \<inter> A j = {}"
 | |
| 70 | by (auto simp: disjoint_family_on_def) | |
| 71 | ||
| 72 | lemma disjoint_family_subset: "disjoint_family A \<Longrightarrow> (\<And>x. B x \<subseteq> A x) \<Longrightarrow> disjoint_family B" | |
| 73 | by (force simp add: disjoint_family_on_def) | |
| 74 | ||
| 75 | lemma disjoint_family_on_bisimulation: | |
| 76 | assumes "disjoint_family_on f S" | |
| 77 |   and "\<And>n m. n \<in> S \<Longrightarrow> m \<in> S \<Longrightarrow> n \<noteq> m \<Longrightarrow> f n \<inter> f m = {} \<Longrightarrow> g n \<inter> g m = {}"
 | |
| 78 | shows "disjoint_family_on g S" | |
| 79 | using assms unfolding disjoint_family_on_def by auto | |
| 80 | ||
| 81 | lemma disjoint_family_on_mono: | |
| 82 | "A \<subseteq> B \<Longrightarrow> disjoint_family_on f B \<Longrightarrow> disjoint_family_on f A" | |
| 83 | unfolding disjoint_family_on_def by auto | |
| 84 | ||
| 85 | lemma disjoint_family_Suc: | |
| 86 | "(\<And>n. A n \<subseteq> A (Suc n)) \<Longrightarrow> disjoint_family (\<lambda>i. A (Suc i) - A i)" | |
| 87 | using lift_Suc_mono_le[of A] | |
| 88 | by (auto simp add: disjoint_family_on_def) | |
| 61824 
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
 paulson <lp15@cam.ac.uk> parents: 
60727diff
changeset | 89 | (metis insert_absorb insert_subset le_SucE le_antisym not_le_imp_less less_imp_le) | 
| 60727 | 90 | |
| 91 | lemma disjoint_family_on_disjoint_image: | |
| 92 | "disjoint_family_on A I \<Longrightarrow> disjoint (A ` I)" | |
| 93 | unfolding disjoint_family_on_def disjoint_def by force | |
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 94 | |
| 60727 | 95 | lemma disjoint_family_on_vimageI: "disjoint_family_on F I \<Longrightarrow> disjoint_family_on (\<lambda>i. f -` F i) I" | 
| 96 | by (auto simp: disjoint_family_on_def) | |
| 97 | ||
| 98 | lemma disjoint_image_disjoint_family_on: | |
| 99 | assumes d: "disjoint (A ` I)" and i: "inj_on A I" | |
| 100 | shows "disjoint_family_on A I" | |
| 101 | unfolding disjoint_family_on_def | |
| 102 | proof (intro ballI impI) | |
| 103 | fix n m assume nm: "m \<in> I" "n \<in> I" and "n \<noteq> m" | |
| 104 |   with i[THEN inj_onD, of n m] show "A n \<inter> A m = {}"
 | |
| 105 | by (intro disjointD[OF d]) auto | |
| 106 | qed | |
| 107 | ||
| 108 | lemma disjoint_UN: | |
| 109 | assumes F: "\<And>i. i \<in> I \<Longrightarrow> disjoint (F i)" and *: "disjoint_family_on (\<lambda>i. \<Union>F i) I" | |
| 110 | shows "disjoint (\<Union>i\<in>I. F i)" | |
| 111 | proof (safe intro!: disjointI del: equalityI) | |
| 112 | fix A B i j assume "A \<noteq> B" "A \<in> F i" "i \<in> I" "B \<in> F j" "j \<in> I" | |
| 113 |   show "A \<inter> B = {}"
 | |
| 114 | proof cases | |
| 115 |     assume "i = j" with F[of i] \<open>i \<in> I\<close> \<open>A \<in> F i\<close> \<open>B \<in> F j\<close> \<open>A \<noteq> B\<close> show "A \<inter> B = {}"
 | |
| 116 | by (auto dest: disjointD) | |
| 117 | next | |
| 118 | assume "i \<noteq> j" | |
| 119 |     with * \<open>i\<in>I\<close> \<open>j\<in>I\<close> have "(\<Union>F i) \<inter> (\<Union>F j) = {}"
 | |
| 120 | by (rule disjoint_family_onD) | |
| 121 | with \<open>A\<in>F i\<close> \<open>i\<in>I\<close> \<open>B\<in>F j\<close> \<open>j\<in>I\<close> | |
| 122 |     show "A \<inter> B = {}"
 | |
| 123 | by auto | |
| 124 | qed | |
| 125 | qed | |
| 126 | ||
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 127 | lemma distinct_list_bind: | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 128 | assumes "distinct xs" "\<And>x. x \<in> set xs \<Longrightarrow> distinct (f x)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 129 | "disjoint_family_on (set \<circ> f) (set xs)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 130 | shows "distinct (List.bind xs f)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 131 | using assms | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 132 | by (induction xs) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 133 | (auto simp: disjoint_family_on_def distinct_map inj_on_def set_list_bind) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 134 | |
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 135 | lemma bij_betw_UNION_disjoint: | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 136 | assumes disj: "disjoint_family_on A' I" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 137 | assumes bij: "\<And>i. i \<in> I \<Longrightarrow> bij_betw f (A i) (A' i)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 138 | shows "bij_betw f (\<Union>i\<in>I. A i) (\<Union>i\<in>I. A' i)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 139 | unfolding bij_betw_def | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 140 | proof | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 141 | from bij show eq: "f ` UNION I A = UNION I A'" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 142 | by (auto simp: bij_betw_def image_UN) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 143 | show "inj_on f (UNION I A)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 144 | proof (rule inj_onI, clarify) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 145 | fix i j x y assume A: "i \<in> I" "j \<in> I" "x \<in> A i" "y \<in> A j" and B: "f x = f y" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 146 | from A bij[of i] bij[of j] have "f x \<in> A' i" "f y \<in> A' j" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 147 | by (auto simp: bij_betw_def) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 148 |     with B have "A' i \<inter> A' j \<noteq> {}" by auto
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 149 | with disj A have "i = j" unfolding disjoint_family_on_def by blast | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 150 | with A B bij[of i] show "x = y" by (auto simp: bij_betw_def dest: inj_onD) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 151 | qed | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 152 | qed | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 153 | |
| 60727 | 154 | lemma disjoint_union: "disjoint C \<Longrightarrow> disjoint B \<Longrightarrow> \<Union>C \<inter> \<Union>B = {} \<Longrightarrow> disjoint (C \<union> B)"
 | 
| 155 |   using disjoint_UN[of "{C, B}" "\<lambda>x. x"] by (auto simp add: disjoint_family_on_def)
 | |
| 156 | ||
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 157 | text \<open> | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 158 | The union of an infinite disjoint family of non-empty sets is infinite. | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 159 | \<close> | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 160 | lemma infinite_disjoint_family_imp_infinite_UNION: | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 161 |   assumes "\<not>finite A" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> {}" "disjoint_family_on f A"
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 162 | shows "\<not>finite (UNION A f)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 163 | proof - | 
| 63148 | 164 | define g where "g x = (SOME y. y \<in> f x)" for x | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 165 | have g: "g x \<in> f x" if "x \<in> A" for x | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 166 | unfolding g_def by (rule someI_ex, insert assms(2) that) blast | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 167 | have inj_on_g: "inj_on g A" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 168 | proof (rule inj_onI, rule ccontr) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 169 | fix x y assume A: "x \<in> A" "y \<in> A" "g x = g y" "x \<noteq> y" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 170 | with g[of x] g[of y] have "g x \<in> f x" "g x \<in> f y" by auto | 
| 63145 | 171 | with A \<open>x \<noteq> y\<close> assms show False | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 172 | by (auto simp: disjoint_family_on_def inj_on_def) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 173 | qed | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 174 | from g have "g ` A \<subseteq> UNION A f" by blast | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 175 | moreover from inj_on_g \<open>\<not>finite A\<close> have "\<not>finite (g ` A)" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 176 | using finite_imageD by blast | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 177 | ultimately show ?thesis using finite_subset by blast | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 178 | qed | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 179 | |
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
62843diff
changeset | 180 | |
| 60727 | 181 | subsection \<open>Construct Disjoint Sequences\<close> | 
| 182 | ||
| 183 | definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set" where | |
| 184 |   "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
 | |
| 185 | ||
| 186 | lemma finite_UN_disjointed_eq: "(\<Union>i\<in>{0..<n}. disjointed A i) = (\<Union>i\<in>{0..<n}. A i)"
 | |
| 187 | proof (induct n) | |
| 188 | case 0 show ?case by simp | |
| 189 | next | |
| 190 | case (Suc n) | |
| 191 | thus ?case by (simp add: atLeastLessThanSuc disjointed_def) | |
| 192 | qed | |
| 193 | ||
| 194 | lemma UN_disjointed_eq: "(\<Union>i. disjointed A i) = (\<Union>i. A i)" | |
| 195 | by (rule UN_finite2_eq [where k=0]) | |
| 196 | (simp add: finite_UN_disjointed_eq) | |
| 197 | ||
| 198 | lemma less_disjoint_disjointed: "m < n \<Longrightarrow> disjointed A m \<inter> disjointed A n = {}"
 | |
| 199 | by (auto simp add: disjointed_def) | |
| 200 | ||
| 201 | lemma disjoint_family_disjointed: "disjoint_family (disjointed A)" | |
| 202 | by (simp add: disjoint_family_on_def) | |
| 203 | (metis neq_iff Int_commute less_disjoint_disjointed) | |
| 204 | ||
| 205 | lemma disjointed_subset: "disjointed A n \<subseteq> A n" | |
| 206 | by (auto simp add: disjointed_def) | |
| 207 | ||
| 208 | lemma disjointed_0[simp]: "disjointed A 0 = A 0" | |
| 209 | by (simp add: disjointed_def) | |
| 210 | ||
| 211 | lemma disjointed_mono: "mono A \<Longrightarrow> disjointed A (Suc n) = A (Suc n) - A n" | |
| 212 | using mono_Un[of A] by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost) | |
| 213 | ||
| 63098 | 214 | subsection \<open>Partitions\<close> | 
| 215 | ||
| 216 | text \<open> | |
| 217 |   Partitions @{term P} of a set @{term A}. We explicitly disallow empty sets.
 | |
| 218 | \<close> | |
| 219 | ||
| 220 | definition partition_on :: "'a set \<Rightarrow> 'a set set \<Rightarrow> bool" | |
| 221 | where | |
| 222 |   "partition_on A P \<longleftrightarrow> \<Union>P = A \<and> disjoint P \<and> {} \<notin> P"
 | |
| 223 | ||
| 224 | lemma partition_onI: | |
| 225 |   "\<Union>P = A \<Longrightarrow> (\<And>p q. p \<in> P \<Longrightarrow> q \<in> P \<Longrightarrow> p \<noteq> q \<Longrightarrow> disjnt p q) \<Longrightarrow> {} \<notin> P \<Longrightarrow> partition_on A P"
 | |
| 226 | by (auto simp: partition_on_def pairwise_def) | |
| 227 | ||
| 228 | lemma partition_onD1: "partition_on A P \<Longrightarrow> A = \<Union>P" | |
| 229 | by (auto simp: partition_on_def) | |
| 230 | ||
| 231 | lemma partition_onD2: "partition_on A P \<Longrightarrow> disjoint P" | |
| 232 | by (auto simp: partition_on_def) | |
| 233 | ||
| 234 | lemma partition_onD3: "partition_on A P \<Longrightarrow> {} \<notin> P"
 | |
| 235 | by (auto simp: partition_on_def) | |
| 236 | ||
| 237 | subsection \<open>Constructions of partitions\<close> | |
| 238 | ||
| 239 | lemma partition_on_empty: "partition_on {} P \<longleftrightarrow> P = {}"
 | |
| 240 | unfolding partition_on_def by fastforce | |
| 241 | ||
| 242 | lemma partition_on_space: "A \<noteq> {} \<Longrightarrow> partition_on A {A}"
 | |
| 243 | by (auto simp: partition_on_def disjoint_def) | |
| 244 | ||
| 245 | lemma partition_on_singletons: "partition_on A ((\<lambda>x. {x}) ` A)"
 | |
| 246 | by (auto simp: partition_on_def disjoint_def) | |
| 247 | ||
| 248 | lemma partition_on_transform: | |
| 249 | assumes P: "partition_on A P" | |
| 250 | assumes F_UN: "\<Union>(F ` P) = F (\<Union>P)" and F_disjnt: "\<And>p q. p \<in> P \<Longrightarrow> q \<in> P \<Longrightarrow> disjnt p q \<Longrightarrow> disjnt (F p) (F q)" | |
| 251 |   shows "partition_on (F A) (F ` P - {{}})"
 | |
| 252 | proof - | |
| 253 |   have "\<Union>(F ` P - {{}}) = F A"
 | |
| 254 | unfolding P[THEN partition_onD1] F_UN[symmetric] by auto | |
| 255 | with P show ?thesis | |
| 256 | by (auto simp add: partition_on_def pairwise_def intro!: F_disjnt) | |
| 257 | qed | |
| 258 | ||
| 259 | lemma partition_on_restrict: "partition_on A P \<Longrightarrow> partition_on (B \<inter> A) (op \<inter> B ` P - {{}})"
 | |
| 260 | by (intro partition_on_transform) (auto simp: disjnt_def) | |
| 261 | ||
| 262 | lemma partition_on_vimage: "partition_on A P \<Longrightarrow> partition_on (f -` A) (op -` f ` P - {{}})"
 | |
| 263 | by (intro partition_on_transform) (auto simp: disjnt_def) | |
| 264 | ||
| 265 | lemma partition_on_inj_image: | |
| 266 | assumes P: "partition_on A P" and f: "inj_on f A" | |
| 267 |   shows "partition_on (f ` A) (op ` f ` P - {{}})"
 | |
| 268 | proof (rule partition_on_transform[OF P]) | |
| 269 | show "p \<in> P \<Longrightarrow> q \<in> P \<Longrightarrow> disjnt p q \<Longrightarrow> disjnt (f ` p) (f ` q)" for p q | |
| 270 | using f[THEN inj_onD] P[THEN partition_onD1] by (auto simp: disjnt_def) | |
| 271 | qed auto | |
| 272 | ||
| 273 | subsection \<open>Finiteness of partitions\<close> | |
| 274 | ||
| 275 | lemma finitely_many_partition_on: | |
| 276 | assumes "finite A" | |
| 277 |   shows "finite {P. partition_on A P}"
 | |
| 278 | proof (rule finite_subset) | |
| 279 |   show "{P. partition_on A P} \<subseteq> Pow (Pow A)"
 | |
| 280 | unfolding partition_on_def by auto | |
| 281 | show "finite (Pow (Pow A))" | |
| 282 | using assms by simp | |
| 283 | qed | |
| 284 | ||
| 285 | lemma finite_elements: "finite A \<Longrightarrow> partition_on A P \<Longrightarrow> finite P" | |
| 286 | using partition_onD1[of A P] by (simp add: finite_UnionD) | |
| 287 | ||
| 288 | subsection \<open>Equivalence of partitions and equivalence classes\<close> | |
| 289 | ||
| 290 | lemma partition_on_quotient: | |
| 291 | assumes r: "equiv A r" | |
| 292 | shows "partition_on A (A // r)" | |
| 293 | proof (rule partition_onI) | |
| 294 | from r have "refl_on A r" | |
| 295 | by (auto elim: equivE) | |
| 296 |   then show "\<Union>(A // r) = A" "{} \<notin> A // r"
 | |
| 297 | by (auto simp: refl_on_def quotient_def) | |
| 298 | ||
| 299 | fix p q assume "p \<in> A // r" "q \<in> A // r" "p \<noteq> q" | |
| 300 |   then obtain x y where "x \<in> A" "y \<in> A" "p = r `` {x}" "q = r `` {y}"
 | |
| 301 | by (auto simp: quotient_def) | |
| 302 | with r equiv_class_eq_iff[OF r, of x y] \<open>p \<noteq> q\<close> show "disjnt p q" | |
| 303 | by (auto simp: disjnt_equiv_class) | |
| 304 | qed | |
| 305 | ||
| 306 | lemma equiv_partition_on: | |
| 307 | assumes P: "partition_on A P" | |
| 308 |   shows "equiv A {(x, y). \<exists>p \<in> P. x \<in> p \<and> y \<in> p}"
 | |
| 309 | proof (rule equivI) | |
| 310 |   have "A = \<Union>P" "disjoint P" "{} \<notin> P"
 | |
| 311 | using P by (auto simp: partition_on_def) | |
| 312 |   then show "refl_on A {(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p}"
 | |
| 313 | unfolding refl_on_def by auto | |
| 314 |   show "trans {(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p}"
 | |
| 315 | using \<open>disjoint P\<close> by (auto simp: trans_def disjoint_def) | |
| 316 | qed (auto simp: sym_def) | |
| 317 | ||
| 318 | lemma partition_on_eq_quotient: | |
| 319 | assumes P: "partition_on A P" | |
| 320 |   shows "A // {(x, y). \<exists>p \<in> P. x \<in> p \<and> y \<in> p} = P"
 | |
| 321 | unfolding quotient_def | |
| 322 | proof safe | |
| 323 | fix x assume "x \<in> A" | |
| 324 | then obtain p where "p \<in> P" "x \<in> p" "\<And>q. q \<in> P \<Longrightarrow> x \<in> q \<Longrightarrow> p = q" | |
| 325 | using P by (auto simp: partition_on_def disjoint_def) | |
| 326 |   then have "{y. \<exists>p\<in>P. x \<in> p \<and> y \<in> p} = p"
 | |
| 327 | by (safe intro!: bexI[of _ p]) simp | |
| 328 |   then show "{(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p} `` {x} \<in> P"
 | |
| 329 | by (simp add: \<open>p \<in> P\<close>) | |
| 330 | next | |
| 331 | fix p assume "p \<in> P" | |
| 332 |   then have "p \<noteq> {}"
 | |
| 333 | using P by (auto simp: partition_on_def) | |
| 334 | then obtain x where "x \<in> p" | |
| 335 | by auto | |
| 336 | then have "x \<in> A" "\<And>q. q \<in> P \<Longrightarrow> x \<in> q \<Longrightarrow> p = q" | |
| 337 | using P \<open>p \<in> P\<close> by (auto simp: partition_on_def disjoint_def) | |
| 338 |   with \<open>p\<in>P\<close> \<open>x \<in> p\<close> have "{y. \<exists>p\<in>P. x \<in> p \<and> y \<in> p} = p"
 | |
| 339 | by (safe intro!: bexI[of _ p]) simp | |
| 340 |   then show "p \<in> (\<Union>x\<in>A. {{(x, y). \<exists>p\<in>P. x \<in> p \<and> y \<in> p} `` {x}})"
 | |
| 341 | by (auto intro: \<open>x \<in> A\<close>) | |
| 342 | qed | |
| 343 | ||
| 344 | lemma partition_on_alt: "partition_on A P \<longleftrightarrow> (\<exists>r. equiv A r \<and> P = A // r)" | |
| 345 | by (auto simp: partition_on_eq_quotient intro!: partition_on_quotient intro: equiv_partition_on) | |
| 346 | ||
| 62390 | 347 | end |