--- a/src/HOL/Equiv_Relations.thy Tue Jan 11 06:47:47 2022 +0000
+++ b/src/HOL/Equiv_Relations.thy Tue Jan 11 06:48:02 2022 +0000
@@ -5,7 +5,7 @@
section \<open>Equivalence Relations in Higher-Order Set Theory\<close>
theory Equiv_Relations
- imports Groups_Big
+ imports BNF_Least_Fixpoint
begin
subsection \<open>Equivalence relations -- set version\<close>
@@ -348,60 +348,6 @@
by (auto simp: Union_quotient dest: quotient_disj)
qed (use assms in blast)
-lemma card_quotient_disjoint:
- assumes "finite A" "inj_on (\<lambda>x. {x} // r) A"
- shows "card (A//r) = card A"
-proof -
- have "\<forall>i\<in>A. \<forall>j\<in>A. i \<noteq> j \<longrightarrow> r `` {j} \<noteq> r `` {i}"
- using assms by (fastforce simp add: quotient_def inj_on_def)
- with assms show ?thesis
- by (simp add: quotient_def card_UN_disjoint)
-qed
-
-text \<open>By Jakub Kądziołka:\<close>
-
-lemma sum_fun_comp:
- assumes "finite S" "finite R" "g ` S \<subseteq> R"
- shows "(\<Sum>x \<in> S. f (g x)) = (\<Sum>y \<in> R. of_nat (card {x \<in> S. g x = y}) * f y)"
-proof -
- let ?r = "relation_of (\<lambda>p q. g p = g q) S"
- have eqv: "equiv S ?r"
- unfolding relation_of_def by (auto intro: comp_equivI)
- have finite: "C \<in> S//?r \<Longrightarrow> finite C" for C
- by (fact finite_equiv_class[OF \<open>finite S\<close> equiv_type[OF \<open>equiv S ?r\<close>]])
- have disjoint: "A \<in> S//?r \<Longrightarrow> B \<in> S//?r \<Longrightarrow> A \<noteq> B \<Longrightarrow> A \<inter> B = {}" for A B
- using eqv quotient_disj by blast
-
- let ?cls = "\<lambda>y. {x \<in> S. y = g x}"
- have quot_as_img: "S//?r = ?cls ` g ` S"
- by (auto simp add: relation_of_def quotient_def)
- have cls_inj: "inj_on ?cls (g ` S)"
- by (auto intro: inj_onI)
-
- have rest_0: "(\<Sum>y \<in> R - g ` S. of_nat (card (?cls y)) * f y) = 0"
- proof -
- have "of_nat (card (?cls y)) * f y = 0" if asm: "y \<in> R - g ` S" for y
- proof -
- from asm have *: "?cls y = {}" by auto
- show ?thesis unfolding * by simp
- qed
- thus ?thesis by simp
- qed
-
- have "(\<Sum>x \<in> S. f (g x)) = (\<Sum>C \<in> S//?r. \<Sum>x \<in> C. f (g x))"
- using eqv finite disjoint
- by (simp flip: sum.Union_disjoint[simplified] add: Union_quotient)
- also have "... = (\<Sum>y \<in> g ` S. \<Sum>x \<in> ?cls y. f (g x))"
- unfolding quot_as_img by (simp add: sum.reindex[OF cls_inj])
- also have "... = (\<Sum>y \<in> g ` S. \<Sum>x \<in> ?cls y. f y)"
- by auto
- also have "... = (\<Sum>y \<in> g ` S. of_nat (card (?cls y)) * f y)"
- by (simp flip: sum_constant)
- also have "... = (\<Sum>y \<in> R. of_nat (card (?cls y)) * f y)"
- using rest_0 by (simp add: sum.subset_diff[OF \<open>g ` S \<subseteq> R\<close> \<open>finite R\<close>])
- finally show ?thesis
- by (simp add: eq_commute)
-qed
subsection \<open>Projection\<close>