# HG changeset patch # User eberlm # Date 1463506062 -7200 # Node ID aa5cffd8a6060d138847aaac13365a94388532e3 # Parent af0e964aad7b6e01513237b9f94509603a280a41# Parent 56f03591898b13d08afd4808ad18aae919f648e9 Merged diff -r af0e964aad7b -r aa5cffd8a606 src/HOL/Library/Disjoint_Sets.thy --- a/src/HOL/Library/Disjoint_Sets.thy Tue May 17 17:05:35 2016 +0200 +++ b/src/HOL/Library/Disjoint_Sets.thy Tue May 17 19:27:42 2016 +0200 @@ -2,7 +2,7 @@ Author: Johannes Hölzl, TU München *) -section \Handling Disjoint Sets\ +section \Partitions and Disjoint Sets\ theory Disjoint_Sets imports Main @@ -20,6 +20,9 @@ lemma mono_Un: "mono A \ (\i\n. A i) = A n" unfolding mono_def by auto +lemma disjnt_equiv_class: "equiv A r \ disjnt (r``{a}) (r``{b}) \ (a, b) \ r" + by (auto dest: equiv_class_self simp: equiv_class_eq_iff disjnt_def) + subsection \Set of Disjoint Sets\ abbreviation disjoint :: "'a set set \ bool" where "disjoint \ pairwise disjnt" @@ -46,7 +49,7 @@ assumes *: "\i. i \ I \ disjoint (F i)" shows "disjoint {\i\I. X i | X. \i\I. X i \ F i}" proof (safe intro!: disjointI del: equalityI) - fix A B :: "'a \ 'b set" assume "(\i\I. A i) \ (\i\I. B i)" + fix A B :: "'a \ 'b set" assume "(\i\I. A i) \ (\i\I. B i)" then obtain i where "A i \ B i" "i \ I" by auto moreover assume "\i\I. A i \ F i" "\i\I. B i \ F i" @@ -207,5 +210,138 @@ lemma disjointed_mono: "mono A \ disjointed A (Suc n) = A (Suc n) - A n" using mono_Un[of A] by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost) - + +subsection \Partitions\ + +text \ + Partitions @{term P} of a set @{term A}. We explicitly disallow empty sets. +\ + +definition partition_on :: "'a set \ 'a set set \ bool" +where + "partition_on A P \ \P = A \ disjoint P \ {} \ P" + +lemma partition_onI: + "\P = A \ (\p q. p \ P \ q \ P \ p \ q \ disjnt p q) \ {} \ P \ partition_on A P" + by (auto simp: partition_on_def pairwise_def) + +lemma partition_onD1: "partition_on A P \ A = \P" + by (auto simp: partition_on_def) + +lemma partition_onD2: "partition_on A P \ disjoint P" + by (auto simp: partition_on_def) + +lemma partition_onD3: "partition_on A P \ {} \ P" + by (auto simp: partition_on_def) + +subsection \Constructions of partitions\ + +lemma partition_on_empty: "partition_on {} P \ P = {}" + unfolding partition_on_def by fastforce + +lemma partition_on_space: "A \ {} \ partition_on A {A}" + by (auto simp: partition_on_def disjoint_def) + +lemma partition_on_singletons: "partition_on A ((\x. {x}) ` A)" + by (auto simp: partition_on_def disjoint_def) + +lemma partition_on_transform: + assumes P: "partition_on A P" + assumes F_UN: "\(F ` P) = F (\P)" and F_disjnt: "\p q. p \ P \ q \ P \ disjnt p q \ disjnt (F p) (F q)" + shows "partition_on (F A) (F ` P - {{}})" +proof - + have "\(F ` P - {{}}) = F A" + unfolding P[THEN partition_onD1] F_UN[symmetric] by auto + with P show ?thesis + by (auto simp add: partition_on_def pairwise_def intro!: F_disjnt) +qed + +lemma partition_on_restrict: "partition_on A P \ partition_on (B \ A) (op \ B ` P - {{}})" + by (intro partition_on_transform) (auto simp: disjnt_def) + +lemma partition_on_vimage: "partition_on A P \ partition_on (f -` A) (op -` f ` P - {{}})" + by (intro partition_on_transform) (auto simp: disjnt_def) + +lemma partition_on_inj_image: + assumes P: "partition_on A P" and f: "inj_on f A" + shows "partition_on (f ` A) (op ` f ` P - {{}})" +proof (rule partition_on_transform[OF P]) + show "p \ P \ q \ P \ disjnt p q \ disjnt (f ` p) (f ` q)" for p q + using f[THEN inj_onD] P[THEN partition_onD1] by (auto simp: disjnt_def) +qed auto + +subsection \Finiteness of partitions\ + +lemma finitely_many_partition_on: + assumes "finite A" + shows "finite {P. partition_on A P}" +proof (rule finite_subset) + show "{P. partition_on A P} \ Pow (Pow A)" + unfolding partition_on_def by auto + show "finite (Pow (Pow A))" + using assms by simp +qed + +lemma finite_elements: "finite A \ partition_on A P \ finite P" + using partition_onD1[of A P] by (simp add: finite_UnionD) + +subsection \Equivalence of partitions and equivalence classes\ + +lemma partition_on_quotient: + assumes r: "equiv A r" + shows "partition_on A (A // r)" +proof (rule partition_onI) + from r have "refl_on A r" + by (auto elim: equivE) + then show "\(A // r) = A" "{} \ A // r" + by (auto simp: refl_on_def quotient_def) + + fix p q assume "p \ A // r" "q \ A // r" "p \ q" + then obtain x y where "x \ A" "y \ A" "p = r `` {x}" "q = r `` {y}" + by (auto simp: quotient_def) + with r equiv_class_eq_iff[OF r, of x y] \p \ q\ show "disjnt p q" + by (auto simp: disjnt_equiv_class) +qed + +lemma equiv_partition_on: + assumes P: "partition_on A P" + shows "equiv A {(x, y). \p \ P. x \ p \ y \ p}" +proof (rule equivI) + have "A = \P" "disjoint P" "{} \ P" + using P by (auto simp: partition_on_def) + then show "refl_on A {(x, y). \p\P. x \ p \ y \ p}" + unfolding refl_on_def by auto + show "trans {(x, y). \p\P. x \ p \ y \ p}" + using \disjoint P\ by (auto simp: trans_def disjoint_def) +qed (auto simp: sym_def) + +lemma partition_on_eq_quotient: + assumes P: "partition_on A P" + shows "A // {(x, y). \p \ P. x \ p \ y \ p} = P" + unfolding quotient_def +proof safe + fix x assume "x \ A" + then obtain p where "p \ P" "x \ p" "\q. q \ P \ x \ q \ p = q" + using P by (auto simp: partition_on_def disjoint_def) + then have "{y. \p\P. x \ p \ y \ p} = p" + by (safe intro!: bexI[of _ p]) simp + then show "{(x, y). \p\P. x \ p \ y \ p} `` {x} \ P" + by (simp add: \p \ P\) +next + fix p assume "p \ P" + then have "p \ {}" + using P by (auto simp: partition_on_def) + then obtain x where "x \ p" + by auto + then have "x \ A" "\q. q \ P \ x \ q \ p = q" + using P \p \ P\ by (auto simp: partition_on_def disjoint_def) + with \p\P\ \x \ p\ have "{y. \p\P. x \ p \ y \ p} = p" + by (safe intro!: bexI[of _ p]) simp + then show "p \ (\x\A. {{(x, y). \p\P. x \ p \ y \ p} `` {x}})" + by (auto intro: \x \ A\) +qed + +lemma partition_on_alt: "partition_on A P \ (\r. equiv A r \ P = A // r)" + by (auto simp: partition_on_eq_quotient intro!: partition_on_quotient intro: equiv_partition_on) + end