diff --git a/src/HOL/Groups_Big.thy b/src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy +++ b/src/HOL/Groups_Big.thy @@ -1,1634 +1,1639 @@ (* Title: HOL/Groups_Big.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Big sum and product over finite (non-empty) sets\ theory Groups_Big imports Power begin subsection \Generic monoid operation over a set\ locale comm_monoid_set = comm_monoid begin subsubsection \Standard sum or product indexed by a finite set\ interpretation comp_fun_commute f by standard (simp add: fun_eq_iff left_commute) interpretation comp?: comp_fun_commute "f \ g" by (fact comp_comp_fun_commute) definition F :: "('b \ 'a) \ 'b set \ 'a" where eq_fold: "F g A = Finite_Set.fold (f \ g) \<^bold>1 A" lemma infinite [simp]: "\ finite A \ F g A = \<^bold>1" by (simp add: eq_fold) lemma empty [simp]: "F g {} = \<^bold>1" by (simp add: eq_fold) lemma insert [simp]: "finite A \ x \ A \ F g (insert x A) = g x \<^bold>* F g A" by (simp add: eq_fold) lemma remove: assumes "finite A" and "x \ A" shows "F g A = g x \<^bold>* F g (A - {x})" proof - from \x \ A\ obtain B where B: "A = insert x B" and "x \ B" by (auto dest: mk_disjoint_insert) moreover from \finite A\ B have "finite B" by simp ultimately show ?thesis by simp qed lemma insert_remove: "finite A \ F g (insert x A) = g x \<^bold>* F g (A - {x})" by (cases "x \ A") (simp_all add: remove insert_absorb) lemma insert_if: "finite A \ F g (insert x A) = (if x \ A then F g A else g x \<^bold>* F g A)" by (cases "x \ A") (simp_all add: insert_absorb) lemma neutral: "\x\A. g x = \<^bold>1 \ F g A = \<^bold>1" by (induct A rule: infinite_finite_induct) simp_all lemma neutral_const [simp]: "F (\_. \<^bold>1) A = \<^bold>1" by (simp add: neutral) lemma union_inter: assumes "finite A" and "finite B" shows "F g (A \ B) \<^bold>* F g (A \ B) = F g A \<^bold>* F g B" \ \The reversed orientation looks more natural, but LOOPS as a simprule!\ using assms proof (induct A) case empty then show ?case by simp next case (insert x A) then show ?case by (auto simp: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute) qed corollary union_inter_neutral: assumes "finite A" and "finite B" and "\x \ A \ B. g x = \<^bold>1" shows "F g (A \ B) = F g A \<^bold>* F g B" using assms by (simp add: union_inter [symmetric] neutral) corollary union_disjoint: assumes "finite A" and "finite B" assumes "A \ B = {}" shows "F g (A \ B) = F g A \<^bold>* F g B" using assms by (simp add: union_inter_neutral) lemma union_diff2: assumes "finite A" and "finite B" shows "F g (A \ B) = F g (A - B) \<^bold>* F g (B - A) \<^bold>* F g (A \ B)" proof - have "A \ B = A - B \ (B - A) \ A \ B" by auto with assms show ?thesis by simp (subst union_disjoint, auto)+ qed lemma subset_diff: assumes "B \ A" and "finite A" shows "F g A = F g (A - B) \<^bold>* F g B" proof - from assms have "finite (A - B)" by auto moreover from assms have "finite B" by (rule finite_subset) moreover from assms have "(A - B) \ B = {}" by auto ultimately have "F g (A - B \ B) = F g (A - B) \<^bold>* F g B" by (rule union_disjoint) moreover from assms have "A \ B = A" by auto ultimately show ?thesis by simp qed lemma Int_Diff: assumes "finite A" shows "F g A = F g (A \ B) \<^bold>* F g (A - B)" by (subst subset_diff [where B = "A - B"]) (auto simp: Diff_Diff_Int assms) lemma setdiff_irrelevant: assumes "finite A" shows "F g (A - {x. g x = z}) = F g A" using assms by (induct A) (simp_all add: insert_Diff_if) lemma not_neutral_contains_not_neutral: assumes "F g A \ \<^bold>1" obtains a where "a \ A" and "g a \ \<^bold>1" proof - from assms have "\a\A. g a \ \<^bold>1" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert a A) then show ?case by fastforce qed with that show thesis by blast qed lemma reindex: assumes "inj_on h A" shows "F g (h ` A) = F (g \ h) A" proof (cases "finite A") case True with assms show ?thesis by (simp add: eq_fold fold_image comp_assoc) next case False with assms have "\ finite (h ` A)" by (blast dest: finite_imageD) with False show ?thesis by simp qed lemma cong [fundef_cong]: assumes "A = B" assumes g_h: "\x. x \ B \ g x = h x" shows "F g A = F h B" using g_h unfolding \A = B\ by (induct B rule: infinite_finite_induct) auto lemma cong_simp [cong]: "\ A = B; \x. x \ B =simp=> g x = h x \ \ F (\x. g x) A = F (\x. h x) B" by (rule cong) (simp_all add: simp_implies_def) lemma reindex_cong: assumes "inj_on l B" assumes "A = l ` B" assumes "\x. x \ B \ g (l x) = h x" shows "F g A = F h B" using assms by (simp add: reindex) +lemma image_eq: + assumes "inj_on g A" + shows "F (\x. x) (g ` A) = F g A" + using assms reindex_cong by fastforce + lemma UNION_disjoint: assumes "finite I" and "\i\I. finite (A i)" and "\i\I. \j\I. i \ j \ A i \ A j = {}" shows "F g (\(A ` I)) = F (\x. F g (A x)) I" using assms proof (induction rule: finite_induct) case (insert i I) then have "\j\I. j \ i" by blast with insert.prems have "A i \ \(A ` I) = {}" by blast with insert show ?case by (simp add: union_disjoint) qed auto lemma Union_disjoint: assumes "\A\C. finite A" "\A\C. \B\C. A \ B \ A \ B = {}" shows "F g (\C) = (F \ F) g C" proof (cases "finite C") case True from UNION_disjoint [OF this assms] show ?thesis by simp next case False then show ?thesis by (auto dest: finite_UnionD intro: infinite) qed lemma distrib: "F (\x. g x \<^bold>* h x) A = F g A \<^bold>* F h A" by (induct A rule: infinite_finite_induct) (simp_all add: assoc commute left_commute) lemma Sigma: assumes "finite A" "\x\A. finite (B x)" shows "F (\x. F (g x) (B x)) A = F (case_prod g) (SIGMA x:A. B x)" unfolding Sigma_def proof (subst UNION_disjoint) show "F (\x. F (g x) (B x)) A = F (\x. F (\(x, y). g x y) (\y\B x. {(x, y)})) A" proof (rule cong [OF refl]) show "F (g x) (B x) = F (\(x, y). g x y) (\y\B x. {(x, y)})" if "x \ A" for x using that assms by (simp add: UNION_disjoint) qed qed (use assms in auto) lemma related: assumes Re: "R \<^bold>1 \<^bold>1" and Rop: "\x1 y1 x2 y2. R x1 x2 \ R y1 y2 \ R (x1 \<^bold>* y1) (x2 \<^bold>* y2)" and fin: "finite S" and R_h_g: "\x\S. R (h x) (g x)" shows "R (F h S) (F g S)" using fin by (rule finite_subset_induct) (use assms in auto) lemma mono_neutral_cong_left: assumes "finite T" and "S \ T" and "\i \ T - S. h i = \<^bold>1" and "\x. x \ S \ g x = h x" shows "F g S = F h T" proof- have eq: "T = S \ (T - S)" using \S \ T\ by blast have d: "S \ (T - S) = {}" using \S \ T\ by blast from \finite T\ \S \ T\ have f: "finite S" "finite (T - S)" by (auto intro: finite_subset) show ?thesis using assms(4) by (simp add: union_disjoint [OF f d, unfolded eq [symmetric]] neutral [OF assms(3)]) qed lemma mono_neutral_cong_right: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ (\x. x \ S \ g x = h x) \ F g T = F h S" by (auto intro!: mono_neutral_cong_left [symmetric]) lemma mono_neutral_left: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ F g S = F g T" by (blast intro: mono_neutral_cong_left) lemma mono_neutral_right: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ F g T = F g S" by (blast intro!: mono_neutral_left [symmetric]) lemma mono_neutral_cong: assumes [simp]: "finite T" "finite S" and *: "\i. i \ T - S \ h i = \<^bold>1" "\i. i \ S - T \ g i = \<^bold>1" and gh: "\x. x \ S \ T \ g x = h x" shows "F g S = F h T" proof- have "F g S = F g (S \ T)" by(rule mono_neutral_right)(auto intro: *) also have "\ = F h (S \ T)" using refl gh by(rule cong) also have "\ = F h T" by(rule mono_neutral_left)(auto intro: *) finally show ?thesis . qed lemma reindex_bij_betw: "bij_betw h S T \ F (\x. g (h x)) S = F g T" by (auto simp: bij_betw_def reindex) lemma reindex_bij_witness: assumes witness: "\a. a \ S \ i (j a) = a" "\a. a \ S \ j a \ T" "\b. b \ T \ j (i b) = b" "\b. b \ T \ i b \ S" assumes eq: "\a. a \ S \ h (j a) = g a" shows "F g S = F h T" proof - have "bij_betw j S T" using bij_betw_byWitness[where A=S and f=j and f'=i and A'=T] witness by auto moreover have "F g S = F (\x. h (j x)) S" by (intro cong) (auto simp: eq) ultimately show ?thesis by (simp add: reindex_bij_betw) qed lemma reindex_bij_betw_not_neutral: assumes fin: "finite S'" "finite T'" assumes bij: "bij_betw h (S - S') (T - T')" assumes nn: "\a. a \ S' \ g (h a) = z" "\b. b \ T' \ g b = z" shows "F (\x. g (h x)) S = F g T" proof - have [simp]: "finite S \ finite T" using bij_betw_finite[OF bij] fin by auto show ?thesis proof (cases "finite S") case True with nn have "F (\x. g (h x)) S = F (\x. g (h x)) (S - S')" by (intro mono_neutral_cong_right) auto also have "\ = F g (T - T')" using bij by (rule reindex_bij_betw) also have "\ = F g T" using nn \finite S\ by (intro mono_neutral_cong_left) auto finally show ?thesis . next case False then show ?thesis by simp qed qed lemma reindex_nontrivial: assumes "finite A" and nz: "\x y. x \ A \ y \ A \ x \ y \ h x = h y \ g (h x) = \<^bold>1" shows "F g (h ` A) = F (g \ h) A" proof (subst reindex_bij_betw_not_neutral [symmetric]) show "bij_betw h (A - {x \ A. (g \ h) x = \<^bold>1}) (h ` A - h ` {x \ A. (g \ h) x = \<^bold>1})" using nz by (auto intro!: inj_onI simp: bij_betw_def) qed (use \finite A\ in auto) lemma reindex_bij_witness_not_neutral: assumes fin: "finite S'" "finite T'" assumes witness: "\a. a \ S - S' \ i (j a) = a" "\a. a \ S - S' \ j a \ T - T'" "\b. b \ T - T' \ j (i b) = b" "\b. b \ T - T' \ i b \ S - S'" assumes nn: "\a. a \ S' \ g a = z" "\b. b \ T' \ h b = z" assumes eq: "\a. a \ S \ h (j a) = g a" shows "F g S = F h T" proof - have bij: "bij_betw j (S - (S' \ S)) (T - (T' \ T))" using witness by (intro bij_betw_byWitness[where f'=i]) auto have F_eq: "F g S = F (\x. h (j x)) S" by (intro cong) (auto simp: eq) show ?thesis unfolding F_eq using fin nn eq by (intro reindex_bij_betw_not_neutral[OF _ _ bij]) auto qed lemma delta_remove: assumes fS: "finite S" shows "F (\k. if k = a then b k else c k) S = (if a \ S then b a \<^bold>* F c (S-{a}) else F c (S-{a}))" proof - let ?f = "(\k. if k = a then b k else c k)" show ?thesis proof (cases "a \ S") case False then have "\k\S. ?f k = c k" by simp with False show ?thesis by simp next case True let ?A = "S - {a}" let ?B = "{a}" from True have eq: "S = ?A \ ?B" by blast have dj: "?A \ ?B = {}" by simp from fS have fAB: "finite ?A" "finite ?B" by auto have "F ?f S = F ?f ?A \<^bold>* F ?f ?B" using union_disjoint [OF fAB dj, of ?f, unfolded eq [symmetric]] by simp with True show ?thesis using comm_monoid_set.remove comm_monoid_set_axioms fS by fastforce qed qed lemma delta [simp]: assumes fS: "finite S" shows "F (\k. if k = a then b k else \<^bold>1) S = (if a \ S then b a else \<^bold>1)" by (simp add: delta_remove [OF assms]) lemma delta' [simp]: assumes fin: "finite S" shows "F (\k. if a = k then b k else \<^bold>1) S = (if a \ S then b a else \<^bold>1)" using delta [OF fin, of a b, symmetric] by (auto intro: cong) lemma If_cases: fixes P :: "'b \ bool" and g h :: "'b \ 'a" assumes fin: "finite A" shows "F (\x. if P x then h x else g x) A = F h (A \ {x. P x}) \<^bold>* F g (A \ - {x. P x})" proof - have a: "A = A \ {x. P x} \ A \ -{x. P x}" "(A \ {x. P x}) \ (A \ -{x. P x}) = {}" by blast+ from fin have f: "finite (A \ {x. P x})" "finite (A \ -{x. P x})" by auto let ?g = "\x. if P x then h x else g x" from union_disjoint [OF f a(2), of ?g] a(1) show ?thesis by (subst (1 2) cong) simp_all qed lemma cartesian_product: "F (\x. F (g x) B) A = F (case_prod g) (A \ B)" proof (cases "A = {} \ B = {}") case True then show ?thesis by auto next case False then have "A \ {}" "B \ {}" by auto show ?thesis proof (cases "finite A \ finite B") case True then show ?thesis by (simp add: Sigma) next case False then consider "infinite A" | "infinite B" by auto then have "infinite (A \ B)" by cases (use \A \ {}\ \B \ {}\ in \auto dest: finite_cartesian_productD1 finite_cartesian_productD2\) then show ?thesis using False by auto qed qed lemma inter_restrict: assumes "finite A" shows "F g (A \ B) = F (\x. if x \ B then g x else \<^bold>1) A" proof - let ?g = "\x. if x \ A \ B then g x else \<^bold>1" have "\i\A - A \ B. (if i \ A \ B then g i else \<^bold>1) = \<^bold>1" by simp moreover have "A \ B \ A" by blast ultimately have "F ?g (A \ B) = F ?g A" using \finite A\ by (intro mono_neutral_left) auto then show ?thesis by simp qed lemma inter_filter: "finite A \ F g {x \ A. P x} = F (\x. if P x then g x else \<^bold>1) A" by (simp add: inter_restrict [symmetric, of A "{x. P x}" g, simplified mem_Collect_eq] Int_def) lemma Union_comp: assumes "\A \ B. finite A" and "\A1 A2 x. A1 \ B \ A2 \ B \ A1 \ A2 \ x \ A1 \ x \ A2 \ g x = \<^bold>1" shows "F g (\B) = (F \ F) g B" using assms proof (induct B rule: infinite_finite_induct) case (infinite A) then have "\ finite (\A)" by (blast dest: finite_UnionD) with infinite show ?case by simp next case empty then show ?case by simp next case (insert A B) then have "finite A" "finite B" "finite (\B)" "A \ B" and "\x\A \ \B. g x = \<^bold>1" and H: "F g (\B) = (F \ F) g B" by auto then have "F g (A \ \B) = F g A \<^bold>* F g (\B)" by (simp add: union_inter_neutral) with \finite B\ \A \ B\ show ?case by (simp add: H) qed lemma swap: "F (\i. F (g i) B) A = F (\j. F (\i. g i j) A) B" unfolding cartesian_product by (rule reindex_bij_witness [where i = "\(i, j). (j, i)" and j = "\(i, j). (j, i)"]) auto lemma swap_restrict: "finite A \ finite B \ F (\x. F (g x) {y. y \ B \ R x y}) A = F (\y. F (\x. g x y) {x. x \ A \ R x y}) B" by (simp add: inter_filter) (rule swap) lemma image_gen: assumes fin: "finite S" shows "F h S = F (\y. F h {x. x \ S \ g x = y}) (g ` S)" proof - have "{y. y\ g`S \ g x = y} = {g x}" if "x \ S" for x using that by auto then have "F h S = F (\x. F (\y. h x) {y. y\ g`S \ g x = y}) S" by simp also have "\ = F (\y. F h {x. x \ S \ g x = y}) (g ` S)" by (rule swap_restrict [OF fin finite_imageI [OF fin]]) finally show ?thesis . qed lemma group: assumes fS: "finite S" and fT: "finite T" and fST: "g ` S \ T" shows "F (\y. F h {x. x \ S \ g x = y}) T = F h S" unfolding image_gen[OF fS, of h g] by (auto intro: neutral mono_neutral_right[OF fT fST]) lemma Plus: fixes A :: "'b set" and B :: "'c set" assumes fin: "finite A" "finite B" shows "F g (A <+> B) = F (g \ Inl) A \<^bold>* F (g \ Inr) B" proof - have "A <+> B = Inl ` A \ Inr ` B" by auto moreover from fin have "finite (Inl ` A)" "finite (Inr ` B)" by auto moreover have "Inl ` A \ Inr ` B = {}" by auto moreover have "inj_on Inl A" "inj_on Inr B" by (auto intro: inj_onI) ultimately show ?thesis using fin by (simp add: union_disjoint reindex) qed lemma same_carrier: assumes "finite C" assumes subset: "A \ C" "B \ C" assumes trivial: "\a. a \ C - A \ g a = \<^bold>1" "\b. b \ C - B \ h b = \<^bold>1" shows "F g A = F h B \ F g C = F h C" proof - have "finite A" and "finite B" and "finite (C - A)" and "finite (C - B)" using \finite C\ subset by (auto elim: finite_subset) from subset have [simp]: "A - (C - A) = A" by auto from subset have [simp]: "B - (C - B) = B" by auto from subset have "C = A \ (C - A)" by auto then have "F g C = F g (A \ (C - A))" by simp also have "\ = F g (A - (C - A)) \<^bold>* F g (C - A - A) \<^bold>* F g (A \ (C - A))" using \finite A\ \finite (C - A)\ by (simp only: union_diff2) finally have *: "F g C = F g A" using trivial by simp from subset have "C = B \ (C - B)" by auto then have "F h C = F h (B \ (C - B))" by simp also have "\ = F h (B - (C - B)) \<^bold>* F h (C - B - B) \<^bold>* F h (B \ (C - B))" using \finite B\ \finite (C - B)\ by (simp only: union_diff2) finally have "F h C = F h B" using trivial by simp with * show ?thesis by simp qed lemma same_carrierI: assumes "finite C" assumes subset: "A \ C" "B \ C" assumes trivial: "\a. a \ C - A \ g a = \<^bold>1" "\b. b \ C - B \ h b = \<^bold>1" assumes "F g C = F h C" shows "F g A = F h B" using assms same_carrier [of C A B] by simp lemma eq_general: assumes B: "\y. y \ B \ \!x. x \ A \ h x = y" and A: "\x. x \ A \ h x \ B \ \(h x) = \ x" shows "F \ A = F \ B" proof - have eq: "B = h ` A" by (auto dest: assms) have h: "inj_on h A" using assms by (blast intro: inj_onI) have "F \ A = F (\ \ h) A" using A by auto also have "\ = F \ B" by (simp add: eq reindex h) finally show ?thesis . qed lemma eq_general_inverses: assumes B: "\y. y \ B \ k y \ A \ h(k y) = y" and A: "\x. x \ A \ h x \ B \ k(h x) = x \ \(h x) = \ x" shows "F \ A = F \ B" by (rule eq_general [where h=h]) (force intro: dest: A B)+ subsubsection \HOL Light variant: sum/product indexed by the non-neutral subset\ text \NB only a subset of the properties above are proved\ definition G :: "['b \ 'a,'b set] \ 'a" where "G p I \ if finite {x \ I. p x \ \<^bold>1} then F p {x \ I. p x \ \<^bold>1} else \<^bold>1" lemma finite_Collect_op: shows "\finite {i \ I. x i \ \<^bold>1}; finite {i \ I. y i \ \<^bold>1}\ \ finite {i \ I. x i \<^bold>* y i \ \<^bold>1}" apply (rule finite_subset [where B = "{i \ I. x i \ \<^bold>1} \ {i \ I. y i \ \<^bold>1}"]) using left_neutral by force+ lemma empty' [simp]: "G p {} = \<^bold>1" by (auto simp: G_def) lemma eq_sum [simp]: "finite I \ G p I = F p I" by (auto simp: G_def intro: mono_neutral_cong_left) lemma insert' [simp]: assumes "finite {x \ I. p x \ \<^bold>1}" shows "G p (insert i I) = (if i \ I then G p I else p i \<^bold>* G p I)" proof - have "{x. x = i \ p x \ \<^bold>1 \ x \ I \ p x \ \<^bold>1} = (if p i = \<^bold>1 then {x \ I. p x \ \<^bold>1} else insert i {x \ I. p x \ \<^bold>1})" by auto then show ?thesis using assms by (simp add: G_def conj_disj_distribR insert_absorb) qed lemma distrib_triv': assumes "finite I" shows "G (\i. g i \<^bold>* h i) I = G g I \<^bold>* G h I" by (simp add: assms local.distrib) lemma non_neutral': "G g {x \ I. g x \ \<^bold>1} = G g I" by (simp add: G_def) lemma distrib': assumes "finite {x \ I. g x \ \<^bold>1}" "finite {x \ I. h x \ \<^bold>1}" shows "G (\i. g i \<^bold>* h i) I = G g I \<^bold>* G h I" proof - have "a \<^bold>* a \ a \ a \ \<^bold>1" for a by auto then have "G (\i. g i \<^bold>* h i) I = G (\i. g i \<^bold>* h i) ({i \ I. g i \ \<^bold>1} \ {i \ I. h i \ \<^bold>1})" using assms by (force simp: G_def finite_Collect_op intro!: mono_neutral_cong) also have "\ = G g I \<^bold>* G h I" proof - have "F g ({i \ I. g i \ \<^bold>1} \ {i \ I. h i \ \<^bold>1}) = G g I" "F h ({i \ I. g i \ \<^bold>1} \ {i \ I. h i \ \<^bold>1}) = G h I" by (auto simp: G_def assms intro: mono_neutral_right) then show ?thesis using assms by (simp add: distrib) qed finally show ?thesis . qed lemma cong': assumes "A = B" assumes g_h: "\x. x \ B \ g x = h x" shows "G g A = G h B" using assms by (auto simp: G_def cong: conj_cong intro: cong) lemma mono_neutral_cong_left': assumes "S \ T" and "\i. i \ T - S \ h i = \<^bold>1" and "\x. x \ S \ g x = h x" shows "G g S = G h T" proof - have *: "{x \ S. g x \ \<^bold>1} = {x \ T. h x \ \<^bold>1}" using assms by (metis DiffI subset_eq) then have "finite {x \ S. g x \ \<^bold>1} = finite {x \ T. h x \ \<^bold>1}" by simp then show ?thesis using assms by (auto simp add: G_def * intro: cong) qed lemma mono_neutral_cong_right': "S \ T \ \i \ T - S. g i = \<^bold>1 \ (\x. x \ S \ g x = h x) \ G g T = G h S" by (auto intro!: mono_neutral_cong_left' [symmetric]) lemma mono_neutral_left': "S \ T \ \i \ T - S. g i = \<^bold>1 \ G g S = G g T" by (blast intro: mono_neutral_cong_left') lemma mono_neutral_right': "S \ T \ \i \ T - S. g i = \<^bold>1 \ G g T = G g S" by (blast intro!: mono_neutral_left' [symmetric]) end subsection \Generalized summation over a set\ context comm_monoid_add begin sublocale sum: comm_monoid_set plus 0 defines sum = sum.F and sum' = sum.G .. abbreviation Sum ("\") where "\ \ sum (\x. x)" end text \Now: lots of fancy syntax. First, \<^term>\sum (\x. e) A\ is written \\x\A. e\.\ syntax (ASCII) "_sum" :: "pttrn \ 'a set \ 'b \ 'b::comm_monoid_add" ("(3SUM (_/:_)./ _)" [0, 51, 10] 10) syntax "_sum" :: "pttrn \ 'a set \ 'b \ 'b::comm_monoid_add" ("(2\(_/\_)./ _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\i\A. b" \ "CONST sum (\i. b) A" text \Instead of \<^term>\\x\{x. P}. e\ we introduce the shorter \\x|P. e\.\ syntax (ASCII) "_qsum" :: "pttrn \ bool \ 'a \ 'a" ("(3SUM _ |/ _./ _)" [0, 0, 10] 10) syntax "_qsum" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0, 0, 10] 10) translations "\x|P. t" => "CONST sum (\x. t) {x. P}" print_translation \ let fun sum_tr' [Abs (x, Tx, t), Const (\<^const_syntax>\Collect\, _) $ Abs (y, Ty, P)] = if x <> y then raise Match else let val x' = Syntax_Trans.mark_bound_body (x, Tx); val t' = subst_bound (x', t); val P' = subst_bound (x', P); in Syntax.const \<^syntax_const>\_qsum\ $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t' end | sum_tr' _ = raise Match; in [(\<^const_syntax>\sum\, K sum_tr')] end \ subsubsection \Properties in more restricted classes of structures\ lemma sum_Un: "finite A \ finite B \ sum f (A \ B) = sum f A + sum f B - sum f (A \ B)" for f :: "'b \ 'a::ab_group_add" by (subst sum.union_inter [symmetric]) (auto simp add: algebra_simps) lemma sum_Un2: assumes "finite (A \ B)" shows "sum f (A \ B) = sum f (A - B) + sum f (B - A) + sum f (A \ B)" proof - have "A \ B = A - B \ (B - A) \ A \ B" by auto with assms show ?thesis by simp (subst sum.union_disjoint, auto)+ qed lemma sum_diff1: fixes f :: "'b \ 'a::ab_group_add" assumes "finite A" shows "sum f (A - {a}) = (if a \ A then sum f A - f a else sum f A)" using assms by induct (auto simp: insert_Diff_if) lemma sum_diff: fixes f :: "'b \ 'a::ab_group_add" assumes "finite A" "B \ A" shows "sum f (A - B) = sum f A - sum f B" proof - from assms(2,1) have "finite B" by (rule finite_subset) from this \B \ A\ show ?thesis proof induct case empty thus ?case by simp next case (insert x F) with \finite A\ \finite B\ show ?case by (simp add: Diff_insert[where a=x and B=F] sum_diff1 insert_absorb) qed qed lemma sum_diff1'_aux: fixes f :: "'a \ 'b::ab_group_add" assumes "finite F" "{i \ I. f i \ 0} \ F" shows "sum' f (I - {i}) = (if i \ I then sum' f I - f i else sum' f I)" using assms proof induct case (insert x F) have 1: "finite {x \ I. f x \ 0} \ finite {x \ I. x \ i \ f x \ 0}" by (erule rev_finite_subset) auto have 2: "finite {x \ I. x \ i \ f x \ 0} \ finite {x \ I. f x \ 0}" apply (drule finite_insert [THEN iffD2]) by (erule rev_finite_subset) auto have 3: "finite {i \ I. f i \ 0}" using finite_subset insert by blast show ?case using insert sum_diff1 [of "{i \ I. f i \ 0}" f i] by (auto simp: sum.G_def 1 2 3 set_diff_eq conj_ac) qed (simp add: sum.G_def) lemma sum_diff1': fixes f :: "'a \ 'b::ab_group_add" assumes "finite {i \ I. f i \ 0}" shows "sum' f (I - {i}) = (if i \ I then sum' f I - f i else sum' f I)" by (rule sum_diff1'_aux [OF assms order_refl]) lemma (in ordered_comm_monoid_add) sum_mono: "(\i. i\K \ f i \ g i) \ (\i\K. f i) \ (\i\K. g i)" by (induct K rule: infinite_finite_induct) (use add_mono in auto) lemma (in strict_ordered_comm_monoid_add) sum_strict_mono: assumes "finite A" "A \ {}" and "\x. x \ A \ f x < g x" shows "sum f A < sum g A" using assms proof (induct rule: finite_ne_induct) case singleton then show ?case by simp next case insert then show ?case by (auto simp: add_strict_mono) qed lemma sum_strict_mono_ex1: fixes f g :: "'i \ 'a::ordered_cancel_comm_monoid_add" assumes "finite A" and "\x\A. f x \ g x" and "\a\A. f a < g a" shows "sum f A < sum g A" proof- from assms(3) obtain a where a: "a \ A" "f a < g a" by blast have "sum f A = sum f ((A - {a}) \ {a})" by(simp add: insert_absorb[OF \a \ A\]) also have "\ = sum f (A - {a}) + sum f {a}" using \finite A\ by(subst sum.union_disjoint) auto also have "sum f (A - {a}) \ sum g (A - {a})" by (rule sum_mono) (simp add: assms(2)) also from a have "sum f {a} < sum g {a}" by simp also have "sum g (A - {a}) + sum g {a} = sum g((A - {a}) \ {a})" using \finite A\ by (subst sum.union_disjoint[symmetric]) auto also have "\ = sum g A" by (simp add: insert_absorb[OF \a \ A\]) finally show ?thesis by (auto simp add: add_right_mono add_strict_left_mono) qed lemma sum_mono_inv: fixes f g :: "'i \ 'a :: ordered_cancel_comm_monoid_add" assumes eq: "sum f I = sum g I" assumes le: "\i. i \ I \ f i \ g i" assumes i: "i \ I" assumes I: "finite I" shows "f i = g i" proof (rule ccontr) assume "\ ?thesis" with le[OF i] have "f i < g i" by simp with i have "\i\I. f i < g i" .. from sum_strict_mono_ex1[OF I _ this] le have "sum f I < sum g I" by blast with eq show False by simp qed lemma member_le_sum: fixes f :: "_ \ 'b::{semiring_1, ordered_comm_monoid_add}" assumes "i \ A" and le: "\x. x \ A - {i} \ 0 \ f x" and "finite A" shows "f i \ sum f A" proof - have "f i \ sum f (A \ {i})" by (simp add: assms) also have "... = (\x\A. if x \ {i} then f x else 0)" using assms sum.inter_restrict by blast also have "... \ sum f A" apply (rule sum_mono) apply (auto simp: le) done finally show ?thesis . qed lemma sum_negf: "(\x\A. - f x) = - (\x\A. f x)" for f :: "'b \ 'a::ab_group_add" by (induct A rule: infinite_finite_induct) auto lemma sum_subtractf: "(\x\A. f x - g x) = (\x\A. f x) - (\x\A. g x)" for f g :: "'b \'a::ab_group_add" using sum.distrib [of f "- g" A] by (simp add: sum_negf) lemma sum_subtractf_nat: "(\x. x \ A \ g x \ f x) \ (\x\A. f x - g x) = (\x\A. f x) - (\x\A. g x)" for f g :: "'a \ nat" by (induct A rule: infinite_finite_induct) (auto simp: sum_mono) context ordered_comm_monoid_add begin lemma sum_nonneg: "(\x. x \ A \ 0 \ f x) \ 0 \ sum f A" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) then have "0 + 0 \ f x + sum f F" by (blast intro: add_mono) with insert show ?case by simp qed lemma sum_nonpos: "(\x. x \ A \ f x \ 0) \ sum f A \ 0" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) then have "f x + sum f F \ 0 + 0" by (blast intro: add_mono) with insert show ?case by simp qed lemma sum_nonneg_eq_0_iff: "finite A \ (\x. x \ A \ 0 \ f x) \ sum f A = 0 \ (\x\A. f x = 0)" by (induct set: finite) (simp_all add: add_nonneg_eq_0_iff sum_nonneg) lemma sum_nonneg_0: "finite s \ (\i. i \ s \ f i \ 0) \ (\ i \ s. f i) = 0 \ i \ s \ f i = 0" by (simp add: sum_nonneg_eq_0_iff) lemma sum_nonneg_leq_bound: assumes "finite s" "\i. i \ s \ f i \ 0" "(\i \ s. f i) = B" "i \ s" shows "f i \ B" proof - from assms have "f i \ f i + (\i \ s - {i}. f i)" by (intro add_increasing2 sum_nonneg) auto also have "\ = B" using sum.remove[of s i f] assms by simp finally show ?thesis by auto qed lemma sum_mono2: assumes fin: "finite B" and sub: "A \ B" and nn: "\b. b \ B-A \ 0 \ f b" shows "sum f A \ sum f B" proof - have "sum f A \ sum f A + sum f (B-A)" by (auto intro: add_increasing2 [OF sum_nonneg] nn) also from fin finite_subset[OF sub fin] have "\ = sum f (A \ (B-A))" by (simp add: sum.union_disjoint del: Un_Diff_cancel) also from sub have "A \ (B-A) = B" by blast finally show ?thesis . qed lemma sum_le_included: assumes "finite s" "finite t" and "\y\t. 0 \ g y" "(\x\s. \y\t. i y = x \ f x \ g y)" shows "sum f s \ sum g t" proof - have "sum f s \ sum (\y. sum g {x. x\t \ i x = y}) s" proof (rule sum_mono) fix y assume "y \ s" with assms obtain z where z: "z \ t" "y = i z" "f y \ g z" by auto with assms show "f y \ sum g {x \ t. i x = y}" (is "?A y \ ?B y") using order_trans[of "?A (i z)" "sum g {z}" "?B (i z)", intro] by (auto intro!: sum_mono2) qed also have "\ \ sum (\y. sum g {x. x\t \ i x = y}) (i ` t)" using assms(2-4) by (auto intro!: sum_mono2 sum_nonneg) also have "\ \ sum g t" using assms by (auto simp: sum.image_gen[symmetric]) finally show ?thesis . qed end lemma (in canonically_ordered_monoid_add) sum_eq_0_iff [simp]: "finite F \ (sum f F = 0) = (\a\F. f a = 0)" by (intro ballI sum_nonneg_eq_0_iff zero_le) context semiring_0 begin lemma sum_distrib_left: "r * sum f A = (\n\A. r * f n)" by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) lemma sum_distrib_right: "sum f A * r = (\n\A. f n * r)" by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) end lemma sum_divide_distrib: "sum f A / r = (\n\A. f n / r)" for r :: "'a::field" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case insert then show ?case by (simp add: add_divide_distrib) qed lemma sum_abs[iff]: "\sum f A\ \ sum (\i. \f i\) A" for f :: "'a \ 'b::ordered_ab_group_add_abs" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case insert then show ?case by (auto intro: abs_triangle_ineq order_trans) qed lemma sum_abs_ge_zero[iff]: "0 \ sum (\i. \f i\) A" for f :: "'a \ 'b::ordered_ab_group_add_abs" by (simp add: sum_nonneg) lemma abs_sum_abs[simp]: "\\a\A. \f a\\ = (\a\A. \f a\)" for f :: "'a \ 'b::ordered_ab_group_add_abs" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert a A) then have "\\a\insert a A. \f a\\ = \\f a\ + (\a\A. \f a\)\" by simp also from insert have "\ = \\f a\ + \\a\A. \f a\\\" by simp also have "\ = \f a\ + \\a\A. \f a\\" by (simp del: abs_of_nonneg) also from insert have "\ = (\a\insert a A. \f a\)" by simp finally show ?case . qed lemma sum_product: fixes f :: "'a \ 'b::semiring_0" shows "sum f A * sum g B = (\i\A. \j\B. f i * g j)" by (simp add: sum_distrib_left sum_distrib_right) (rule sum.swap) lemma sum_mult_sum_if_inj: fixes f :: "'a \ 'b::semiring_0" shows "inj_on (\(a, b). f a * g b) (A \ B) \ sum f A * sum g B = sum id {f a * g b |a b. a \ A \ b \ B}" by(auto simp: sum_product sum.cartesian_product intro!: sum.reindex_cong[symmetric]) lemma sum_SucD: "sum f A = Suc n \ \a\A. 0 < f a" by (induct A rule: infinite_finite_induct) auto lemma sum_eq_Suc0_iff: "finite A \ sum f A = Suc 0 \ (\a\A. f a = Suc 0 \ (\b\A. a \ b \ f b = 0))" by (induct A rule: finite_induct) (auto simp add: add_is_1) lemmas sum_eq_1_iff = sum_eq_Suc0_iff[simplified One_nat_def[symmetric]] lemma sum_Un_nat: "finite A \ finite B \ sum f (A \ B) = sum f A + sum f B - sum f (A \ B)" for f :: "'a \ nat" \ \For the natural numbers, we have subtraction.\ by (subst sum.union_inter [symmetric]) (auto simp: algebra_simps) lemma sum_diff1_nat: "sum f (A - {a}) = (if a \ A then sum f A - f a else sum f A)" for f :: "'a \ nat" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case insert then show ?case apply (auto simp: insert_Diff_if) apply (drule mk_disjoint_insert) apply auto done qed lemma sum_diff_nat: fixes f :: "'a \ nat" assumes "finite B" and "B \ A" shows "sum f (A - B) = sum f A - sum f B" using assms proof induct case empty then show ?case by simp next case (insert x F) note IH = \F \ A \ sum f (A - F) = sum f A - sum f F\ from \x \ F\ \insert x F \ A\ have "x \ A - F" by simp then have A: "sum f ((A - F) - {x}) = sum f (A - F) - f x" by (simp add: sum_diff1_nat) from \insert x F \ A\ have "F \ A" by simp with IH have "sum f (A - F) = sum f A - sum f F" by simp with A have B: "sum f ((A - F) - {x}) = sum f A - sum f F - f x" by simp from \x \ F\ have "A - insert x F = (A - F) - {x}" by auto with B have C: "sum f (A - insert x F) = sum f A - sum f F - f x" by simp from \finite F\ \x \ F\ have "sum f (insert x F) = sum f F + f x" by simp with C have "sum f (A - insert x F) = sum f A - sum f (insert x F)" by simp then show ?case by simp qed lemma sum_comp_morphism: "h 0 = 0 \ (\x y. h (x + y) = h x + h y) \ sum (h \ g) A = h (sum g A)" by (induct A rule: infinite_finite_induct) simp_all lemma (in comm_semiring_1) dvd_sum: "(\a. a \ A \ d dvd f a) \ d dvd sum f A" by (induct A rule: infinite_finite_induct) simp_all lemma (in ordered_comm_monoid_add) sum_pos: "finite I \ I \ {} \ (\i. i \ I \ 0 < f i) \ 0 < sum f I" by (induct I rule: finite_ne_induct) (auto intro: add_pos_pos) lemma (in ordered_comm_monoid_add) sum_pos2: assumes I: "finite I" "i \ I" "0 < f i" "\i. i \ I \ 0 \ f i" shows "0 < sum f I" proof - have "0 < f i + sum f (I - {i})" using assms by (intro add_pos_nonneg sum_nonneg) auto also have "\ = sum f I" using assms by (simp add: sum.remove) finally show ?thesis . qed lemma sum_cong_Suc: assumes "0 \ A" "\x. Suc x \ A \ f (Suc x) = g (Suc x)" shows "sum f A = sum g A" proof (rule sum.cong) fix x assume "x \ A" with assms(1) show "f x = g x" by (cases x) (auto intro!: assms(2)) qed simp_all subsubsection \Cardinality as special case of \<^const>\sum\\ lemma card_eq_sum: "card A = sum (\x. 1) A" proof - have "plus \ (\_. Suc 0) = (\_. Suc)" by (simp add: fun_eq_iff) then have "Finite_Set.fold (plus \ (\_. Suc 0)) = Finite_Set.fold (\_. Suc)" by (rule arg_cong) then have "Finite_Set.fold (plus \ (\_. Suc 0)) 0 A = Finite_Set.fold (\_. Suc) 0 A" by (blast intro: fun_cong) then show ?thesis by (simp add: card.eq_fold sum.eq_fold) qed context semiring_1 begin lemma sum_constant [simp]: "(\x \ A. y) = of_nat (card A) * y" by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) end lemma sum_Suc: "sum (\x. Suc(f x)) A = sum f A + card A" using sum.distrib[of f "\_. 1" A] by simp lemma sum_bounded_above: fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}" assumes le: "\i. i\A \ f i \ K" shows "sum f A \ of_nat (card A) * K" proof (cases "finite A") case True then show ?thesis using le sum_mono[where K=A and g = "\x. K"] by simp next case False then show ?thesis by simp qed lemma sum_bounded_above_divide: fixes K :: "'a::linordered_field" assumes le: "\i. i\A \ f i \ K / of_nat (card A)" and fin: "finite A" "A \ {}" shows "sum f A \ K" using sum_bounded_above [of A f "K / of_nat (card A)", OF le] fin by simp lemma sum_bounded_above_strict: fixes K :: "'a::{ordered_cancel_comm_monoid_add,semiring_1}" assumes "\i. i\A \ f i < K" "card A > 0" shows "sum f A < of_nat (card A) * K" using assms sum_strict_mono[where A=A and g = "\x. K"] by (simp add: card_gt_0_iff) lemma sum_bounded_below: fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}" assumes le: "\i. i\A \ K \ f i" shows "of_nat (card A) * K \ sum f A" proof (cases "finite A") case True then show ?thesis using le sum_mono[where K=A and f = "\x. K"] by simp next case False then show ?thesis by simp qed lemma convex_sum_bound_le: fixes x :: "'a \ 'b::linordered_idom" assumes 0: "\i. i \ I \ 0 \ x i" and 1: "sum x I = 1" and \: "\i. i \ I \ \a i - b\ \ \" shows "\(\i\I. a i * x i) - b\ \ \" proof - have [simp]: "(\i\I. c * x i) = c" for c by (simp flip: sum_distrib_left 1) then have "\(\i\I. a i * x i) - b\ = \\i\I. (a i - b) * x i\" by (simp add: sum_subtractf left_diff_distrib) also have "\ \ (\i\I. \(a i - b) * x i\)" using abs_abs abs_of_nonneg by blast also have "\ \ (\i\I. \(a i - b)\ * x i)" by (simp add: abs_mult 0) also have "\ \ (\i\I. \ * x i)" by (rule sum_mono) (use \ "0" mult_right_mono in blast) also have "\ = \" by simp finally show ?thesis . qed lemma card_UN_disjoint: assumes "finite I" and "\i\I. finite (A i)" and "\i\I. \j\I. i \ j \ A i \ A j = {}" shows "card (\(A ` I)) = (\i\I. card(A i))" proof - have "(\i\I. card (A i)) = (\i\I. \x\A i. 1)" by simp with assms show ?thesis by (simp add: card_eq_sum sum.UNION_disjoint del: sum_constant) qed lemma card_Union_disjoint: assumes "pairwise disjnt C" and fin: "\A. A \ C \ finite A" shows "card (\C) = sum card C" proof (cases "finite C") case True then show ?thesis using card_UN_disjoint [OF True, of "\x. x"] assms by (simp add: disjnt_def fin pairwise_def) next case False then show ?thesis using assms card_eq_0_iff finite_UnionD by fastforce qed lemma card_Union_le_sum_card: fixes U :: "'a set set" assumes "\u \ U. finite u" shows "card (\U) \ sum card U" proof (cases "finite U") case False then show "card (\U) \ sum card U" using card_eq_0_iff finite_UnionD by auto next case True then show "card (\U) \ sum card U" proof (induct U rule: finite_induct) case empty then show ?case by auto next case (insert x F) then have "card(\(insert x F)) \ card(x) + card (\F)" using card_Un_le by auto also have "... \ card(x) + sum card F" using insert.hyps by auto also have "... = sum card (insert x F)" using sum.insert_if and insert.hyps by auto finally show ?case . qed qed lemma card_UN_le: assumes "finite I" shows "card(\i\I. A i) \ (\i\I. card(A i))" using assms proof induction case (insert i I) then show ?case using card_Un_le nat_add_left_cancel_le by (force intro: order_trans) qed auto lemma sum_multicount_gen: assumes "finite s" "finite t" "\j\t. (card {i\s. R i j} = k j)" shows "sum (\i. (card {j\t. R i j})) s = sum k t" (is "?l = ?r") proof- have "?l = sum (\i. sum (\x.1) {j\t. R i j}) s" by auto also have "\ = ?r" unfolding sum.swap_restrict [OF assms(1-2)] using assms(3) by auto finally show ?thesis . qed lemma sum_multicount: assumes "finite S" "finite T" "\j\T. (card {i\S. R i j} = k)" shows "sum (\i. card {j\T. R i j}) S = k * card T" (is "?l = ?r") proof- have "?l = sum (\i. k) T" by (rule sum_multicount_gen) (auto simp: assms) also have "\ = ?r" by (simp add: mult.commute) finally show ?thesis by auto qed lemma sum_card_image: assumes "finite A" assumes "pairwise (\s t. disjnt (f s) (f t)) A" shows "sum card (f ` A) = sum (\a. card (f a)) A" using assms proof (induct A) case (insert a A) show ?case proof cases assume "f a = {}" with insert show ?case by (subst sum.mono_neutral_right[where S="f ` A"]) (auto simp: pairwise_insert) next assume "f a \ {}" then have "sum card (insert (f a) (f ` A)) = card (f a) + sum card (f ` A)" using insert by (subst sum.insert) (auto simp: pairwise_insert) with insert show ?case by (simp add: pairwise_insert) qed qed simp subsubsection \Cardinality of products\ lemma card_SigmaI [simp]: "finite A \ \a\A. finite (B a) \ card (SIGMA x: A. B x) = (\a\A. card (B a))" by (simp add: card_eq_sum sum.Sigma del: sum_constant) (* lemma SigmaI_insert: "y \ A ==> (SIGMA x:(insert y A). B x) = (({y} \ (B y)) \ (SIGMA x: A. B x))" by auto *) lemma card_cartesian_product: "card (A \ B) = card A * card B" by (cases "finite A \ finite B") (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2) lemma card_cartesian_product_singleton: "card ({x} \ A) = card A" by (simp add: card_cartesian_product) subsection \Generalized product over a set\ context comm_monoid_mult begin sublocale prod: comm_monoid_set times 1 defines prod = prod.F and prod' = prod.G .. abbreviation Prod ("\_" [1000] 999) where "\A \ prod (\x. x) A" end syntax (ASCII) "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(4PROD (_/:_)./ _)" [0, 51, 10] 10) syntax "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(2\(_/\_)./ _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\i\A. b" == "CONST prod (\i. b) A" text \Instead of \<^term>\\x\{x. P}. e\ we introduce the shorter \\x|P. e\.\ syntax (ASCII) "_qprod" :: "pttrn \ bool \ 'a \ 'a" ("(4PROD _ |/ _./ _)" [0, 0, 10] 10) syntax "_qprod" :: "pttrn \ bool \ 'a \ 'a" ("(2\_ | (_)./ _)" [0, 0, 10] 10) translations "\x|P. t" => "CONST prod (\x. t) {x. P}" context comm_monoid_mult begin lemma prod_dvd_prod: "(\a. a \ A \ f a dvd g a) \ prod f A dvd prod g A" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by (auto intro: dvdI) next case empty then show ?case by (auto intro: dvdI) next case (insert a A) then have "f a dvd g a" and "prod f A dvd prod g A" by simp_all then obtain r s where "g a = f a * r" and "prod g A = prod f A * s" by (auto elim!: dvdE) then have "g a * prod g A = f a * prod f A * (r * s)" by (simp add: ac_simps) with insert.hyps show ?case by (auto intro: dvdI) qed lemma prod_dvd_prod_subset: "finite B \ A \ B \ prod f A dvd prod f B" by (auto simp add: prod.subset_diff ac_simps intro: dvdI) end subsubsection \Properties in more restricted classes of structures\ context linordered_nonzero_semiring begin lemma prod_ge_1: "(\x. x \ A \ 1 \ f x) \ 1 \ prod f A" proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) have "1 * 1 \ f x * prod f F" by (rule mult_mono') (use insert in auto) with insert show ?case by simp qed lemma prod_le_1: fixes f :: "'b \ 'a" assumes "\x. x \ A \ 0 \ f x \ f x \ 1" shows "prod f A \ 1" using assms proof (induct A rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert x F) then show ?case by (force simp: mult.commute intro: dest: mult_le_one) qed end context comm_semiring_1 begin lemma dvd_prod_eqI [intro]: assumes "finite A" and "a \ A" and "b = f a" shows "b dvd prod f A" proof - from \finite A\ have "prod f (insert a (A - {a})) = f a * prod f (A - {a})" by (intro prod.insert) auto also from \a \ A\ have "insert a (A - {a}) = A" by blast finally have "prod f A = f a * prod f (A - {a})" . with \b = f a\ show ?thesis by simp qed lemma dvd_prodI [intro]: "finite A \ a \ A \ f a dvd prod f A" by auto lemma prod_zero: assumes "finite A" and "\a\A. f a = 0" shows "prod f A = 0" using assms proof (induct A) case empty then show ?case by simp next case (insert a A) then have "f a = 0 \ (\a\A. f a = 0)" by simp then have "f a * prod f A = 0" by rule (simp_all add: insert) with insert show ?case by simp qed lemma prod_dvd_prod_subset2: assumes "finite B" and "A \ B" and "\a. a \ A \ f a dvd g a" shows "prod f A dvd prod g B" proof - from assms have "prod f A dvd prod g A" by (auto intro: prod_dvd_prod) moreover from assms have "prod g A dvd prod g B" by (auto intro: prod_dvd_prod_subset) ultimately show ?thesis by (rule dvd_trans) qed end lemma (in semidom) prod_zero_iff [simp]: fixes f :: "'b \ 'a" assumes "finite A" shows "prod f A = 0 \ (\a\A. f a = 0)" using assms by (induct A) (auto simp: no_zero_divisors) lemma (in semidom_divide) prod_diff1: assumes "finite A" and "f a \ 0" shows "prod f (A - {a}) = (if a \ A then prod f A div f a else prod f A)" proof (cases "a \ A") case True then show ?thesis by simp next case False with assms show ?thesis proof induct case empty then show ?case by simp next case (insert b B) then show ?case proof (cases "a = b") case True with insert show ?thesis by simp next case False with insert have "a \ B" by simp define C where "C = B - {a}" with \finite B\ \a \ B\ have "B = insert a C" "finite C" "a \ C" by auto with insert show ?thesis by (auto simp add: insert_commute ac_simps) qed qed qed lemma sum_zero_power [simp]: "(\i\A. c i * 0^i) = (if finite A \ 0 \ A then c 0 else 0)" for c :: "nat \ 'a::division_ring" by (induct A rule: infinite_finite_induct) auto lemma sum_zero_power' [simp]: "(\i\A. c i * 0^i / d i) = (if finite A \ 0 \ A then c 0 / d 0 else 0)" for c :: "nat \ 'a::field" using sum_zero_power [of "\i. c i / d i" A] by auto lemma (in field) prod_inversef: "prod (inverse \ f) A = inverse (prod f A)" proof (cases "finite A") case True then show ?thesis by (induct A rule: finite_induct) simp_all next case False then show ?thesis by auto qed lemma (in field) prod_dividef: "(\x\A. f x / g x) = prod f A / prod g A" using prod_inversef [of g A] by (simp add: divide_inverse prod.distrib) lemma prod_Un: fixes f :: "'b \ 'a :: field" assumes "finite A" and "finite B" and "\x\A \ B. f x \ 0" shows "prod f (A \ B) = prod f A * prod f B / prod f (A \ B)" proof - from assms have "prod f A * prod f B = prod f (A \ B) * prod f (A \ B)" by (simp add: prod.union_inter [symmetric, of A B]) with assms show ?thesis by simp qed context linordered_semidom begin lemma prod_nonneg: "(\a\A. 0 \ f a) \ 0 \ prod f A" by (induct A rule: infinite_finite_induct) simp_all lemma prod_pos: "(\a\A. 0 < f a) \ 0 < prod f A" by (induct A rule: infinite_finite_induct) simp_all lemma prod_mono: "(\i. i \ A \ 0 \ f i \ f i \ g i) \ prod f A \ prod g A" by (induct A rule: infinite_finite_induct) (force intro!: prod_nonneg mult_mono)+ lemma prod_mono_strict: assumes "finite A" "\i. i \ A \ 0 \ f i \ f i < g i" "A \ {}" shows "prod f A < prod g A" using assms proof (induct A rule: finite_induct) case empty then show ?case by simp next case insert then show ?case by (force intro: mult_strict_mono' prod_nonneg) qed end lemma prod_mono2: fixes f :: "'a \ 'b :: linordered_idom" assumes fin: "finite B" and sub: "A \ B" and nn: "\b. b \ B-A \ 1 \ f b" and A: "\a. a \ A \ 0 \ f a" shows "prod f A \ prod f B" proof - have "prod f A \ prod f A * prod f (B-A)" by (metis prod_ge_1 A mult_le_cancel_left1 nn not_less prod_nonneg) also from fin finite_subset[OF sub fin] have "\ = prod f (A \ (B-A))" by (simp add: prod.union_disjoint del: Un_Diff_cancel) also from sub have "A \ (B-A) = B" by blast finally show ?thesis . qed lemma less_1_prod: fixes f :: "'a \ 'b::linordered_idom" shows "finite I \ I \ {} \ (\i. i \ I \ 1 < f i) \ 1 < prod f I" by (induct I rule: finite_ne_induct) (auto intro: less_1_mult) lemma less_1_prod2: fixes f :: "'a \ 'b::linordered_idom" assumes I: "finite I" "i \ I" "1 < f i" "\i. i \ I \ 1 \ f i" shows "1 < prod f I" proof - have "1 < f i * prod f (I - {i})" using assms by (meson DiffD1 leI less_1_mult less_le_trans mult_le_cancel_left1 prod_ge_1) also have "\ = prod f I" using assms by (simp add: prod.remove) finally show ?thesis . qed lemma (in linordered_field) abs_prod: "\prod f A\ = (\x\A. \f x\)" by (induct A rule: infinite_finite_induct) (simp_all add: abs_mult) lemma prod_eq_1_iff [simp]: "finite A \ prod f A = 1 \ (\a\A. f a = 1)" for f :: "'a \ nat" by (induct A rule: finite_induct) simp_all lemma prod_pos_nat_iff [simp]: "finite A \ prod f A > 0 \ (\a\A. f a > 0)" for f :: "'a \ nat" using prod_zero_iff by (simp del: neq0_conv add: zero_less_iff_neq_zero) lemma prod_constant [simp]: "(\x\ A. y) = y ^ card A" for y :: "'a::comm_monoid_mult" by (induct A rule: infinite_finite_induct) simp_all lemma prod_power_distrib: "prod f A ^ n = prod (\x. (f x) ^ n) A" for f :: "'a \ 'b::comm_semiring_1" by (induct A rule: infinite_finite_induct) (auto simp add: power_mult_distrib) lemma power_sum: "c ^ (\a\A. f a) = (\a\A. c ^ f a)" by (induct A rule: infinite_finite_induct) (simp_all add: power_add) lemma prod_gen_delta: fixes b :: "'b \ 'a::comm_monoid_mult" assumes fin: "finite S" shows "prod (\k. if k = a then b k else c) S = (if a \ S then b a * c ^ (card S - 1) else c ^ card S)" proof - let ?f = "(\k. if k=a then b k else c)" show ?thesis proof (cases "a \ S") case False then have "\ k\ S. ?f k = c" by simp with False show ?thesis by (simp add: prod_constant) next case True let ?A = "S - {a}" let ?B = "{a}" from True have eq: "S = ?A \ ?B" by blast have disjoint: "?A \ ?B = {}" by simp from fin have fin': "finite ?A" "finite ?B" by auto have f_A0: "prod ?f ?A = prod (\i. c) ?A" by (rule prod.cong) auto from fin True have card_A: "card ?A = card S - 1" by auto have f_A1: "prod ?f ?A = c ^ card ?A" unfolding f_A0 by (rule prod_constant) have "prod ?f ?A * prod ?f ?B = prod ?f S" using prod.union_disjoint[OF fin' disjoint, of ?f, unfolded eq[symmetric]] by simp with True card_A show ?thesis by (simp add: f_A1 field_simps cong add: prod.cong cong del: if_weak_cong) qed qed lemma sum_image_le: fixes g :: "'a \ 'b::ordered_comm_monoid_add" assumes "finite I" "\i. i \ I \ 0 \ g(f i)" shows "sum g (f ` I) \ sum (g \ f) I" using assms proof induction case empty then show ?case by auto next case (insert x F) from insertI1 have "0 \ g (f x)" by (rule insert) hence 1: "sum g (f ` F) \ g (f x) + sum g (f ` F)" using add_increasing by blast have 2: "sum g (f ` F) \ sum (g \ f) F" using insert by blast have "sum g (f ` insert x F) = sum g (insert (f x) (f ` F))" by simp also have "\ \ g (f x) + sum g (f ` F)" by (simp add: 1 insert sum.insert_if) also from 2 have "\ \ g (f x) + sum (g \ f) F" by (rule add_left_mono) also from insert(1, 2) have "\ = sum (g \ f) (insert x F)" by (simp add: sum.insert_if) finally show ?case . qed end diff --git a/src/HOL/Library/Infinite_Set.thy b/src/HOL/Library/Infinite_Set.thy --- a/src/HOL/Library/Infinite_Set.thy +++ b/src/HOL/Library/Infinite_Set.thy @@ -1,588 +1,590 @@ (* Title: HOL/Library/Infinite_Set.thy Author: Stephan Merz *) section \Infinite Sets and Related Concepts\ theory Infinite_Set imports Main begin subsection \The set of natural numbers is infinite\ lemma infinite_nat_iff_unbounded_le: "infinite S \ (\m. \n\m. n \ S)" for S :: "nat set" using frequently_cofinite[of "\x. x \ S"] by (simp add: cofinite_eq_sequentially frequently_def eventually_sequentially) lemma infinite_nat_iff_unbounded: "infinite S \ (\m. \n>m. n \ S)" for S :: "nat set" using frequently_cofinite[of "\x. x \ S"] by (simp add: cofinite_eq_sequentially frequently_def eventually_at_top_dense) lemma finite_nat_iff_bounded: "finite S \ (\k. S \ {.. (\k. S \ {.. k})" for S :: "nat set" using infinite_nat_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) lemma finite_nat_bounded: "finite S \ \k. S \ {.. For a set of natural numbers to be infinite, it is enough to know that for any number larger than some \k\, there is some larger number that is an element of the set. \ lemma unbounded_k_infinite: "\m>k. \n>m. n \ S \ infinite (S::nat set)" apply (clarsimp simp add: finite_nat_set_iff_bounded) apply (drule_tac x="Suc (max m k)" in spec) using less_Suc_eq apply fastforce done lemma nat_not_finite: "finite (UNIV::nat set) \ R" by simp lemma range_inj_infinite: fixes f :: "nat \ 'a" assumes "inj f" shows "infinite (range f)" proof assume "finite (range f)" from this assms have "finite (UNIV::nat set)" by (rule finite_imageD) then show False by simp qed subsection \The set of integers is also infinite\ lemma infinite_int_iff_infinite_nat_abs: "infinite S \ infinite ((nat \ abs) ` S)" for S :: "int set" proof (unfold Not_eq_iff, rule iffI) assume "finite ((nat \ abs) ` S)" then have "finite (nat ` (abs ` S))" by (simp add: image_image cong: image_cong) moreover have "inj_on nat (abs ` S)" by (rule inj_onI) auto ultimately have "finite (abs ` S)" by (rule finite_imageD) then show "finite S" by (rule finite_image_absD) qed simp proposition infinite_int_iff_unbounded_le: "infinite S \ (\m. \n. \n\ \ m \ n \ S)" for S :: "int set" by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded_le o_def image_def) (metis abs_ge_zero nat_le_eq_zle le_nat_iff) proposition infinite_int_iff_unbounded: "infinite S \ (\m. \n. \n\ > m \ n \ S)" for S :: "int set" by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded o_def image_def) (metis (full_types) nat_le_iff nat_mono not_le) proposition finite_int_iff_bounded: "finite S \ (\k. abs ` S \ {.. (\k. abs ` S \ {.. k})" for S :: "int set" using infinite_int_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) subsection \Infinitely Many and Almost All\ text \ We often need to reason about the existence of infinitely many (resp., all but finitely many) objects satisfying some predicate, so we introduce corresponding binders and their proof rules. \ lemma not_INFM [simp]: "\ (INFM x. P x) \ (MOST x. \ P x)" by (rule not_frequently) lemma not_MOST [simp]: "\ (MOST x. P x) \ (INFM x. \ P x)" by (rule not_eventually) lemma INFM_const [simp]: "(INFM x::'a. P) \ P \ infinite (UNIV::'a set)" by (simp add: frequently_const_iff) lemma MOST_const [simp]: "(MOST x::'a. P) \ P \ finite (UNIV::'a set)" by (simp add: eventually_const_iff) lemma INFM_imp_distrib: "(INFM x. P x \ Q x) \ ((MOST x. P x) \ (INFM x. Q x))" by (rule frequently_imp_iff) lemma MOST_imp_iff: "MOST x. P x \ (MOST x. P x \ Q x) \ (MOST x. Q x)" by (auto intro: eventually_rev_mp eventually_mono) lemma INFM_conjI: "INFM x. P x \ MOST x. Q x \ INFM x. P x \ Q x" by (rule frequently_rev_mp[of P]) (auto elim: eventually_mono) text \Properties of quantifiers with injective functions.\ lemma INFM_inj: "INFM x. P (f x) \ inj f \ INFM x. P x" using finite_vimageI[of "{x. P x}" f] by (auto simp: frequently_cofinite) lemma MOST_inj: "MOST x. P x \ inj f \ MOST x. P (f x)" using finite_vimageI[of "{x. \ P x}" f] by (auto simp: eventually_cofinite) text \Properties of quantifiers with singletons.\ lemma not_INFM_eq [simp]: "\ (INFM x. x = a)" "\ (INFM x. a = x)" unfolding frequently_cofinite by simp_all lemma MOST_neq [simp]: "MOST x. x \ a" "MOST x. a \ x" unfolding eventually_cofinite by simp_all lemma INFM_neq [simp]: "(INFM x::'a. x \ a) \ infinite (UNIV::'a set)" "(INFM x::'a. a \ x) \ infinite (UNIV::'a set)" unfolding frequently_cofinite by simp_all lemma MOST_eq [simp]: "(MOST x::'a. x = a) \ finite (UNIV::'a set)" "(MOST x::'a. a = x) \ finite (UNIV::'a set)" unfolding eventually_cofinite by simp_all lemma MOST_eq_imp: "MOST x. x = a \ P x" "MOST x. a = x \ P x" unfolding eventually_cofinite by simp_all text \Properties of quantifiers over the naturals.\ lemma MOST_nat: "(\\<^sub>\n. P n) \ (\m. \n>m. P n)" for P :: "nat \ bool" by (auto simp add: eventually_cofinite finite_nat_iff_bounded_le subset_eq simp flip: not_le) lemma MOST_nat_le: "(\\<^sub>\n. P n) \ (\m. \n\m. P n)" for P :: "nat \ bool" by (auto simp add: eventually_cofinite finite_nat_iff_bounded subset_eq simp flip: not_le) lemma INFM_nat: "(\\<^sub>\n. P n) \ (\m. \n>m. P n)" for P :: "nat \ bool" by (simp add: frequently_cofinite infinite_nat_iff_unbounded) lemma INFM_nat_le: "(\\<^sub>\n. P n) \ (\m. \n\m. P n)" for P :: "nat \ bool" by (simp add: frequently_cofinite infinite_nat_iff_unbounded_le) lemma MOST_INFM: "infinite (UNIV::'a set) \ MOST x::'a. P x \ INFM x::'a. P x" by (simp add: eventually_frequently) lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \ (MOST n. P n)" by (simp add: cofinite_eq_sequentially) lemma MOST_SucI: "MOST n. P n \ MOST n. P (Suc n)" and MOST_SucD: "MOST n. P (Suc n) \ MOST n. P n" by (simp_all add: MOST_Suc_iff) lemma MOST_ge_nat: "MOST n::nat. m \ n" by (simp add: cofinite_eq_sequentially) \ \legacy names\ lemma Inf_many_def: "Inf_many P \ infinite {x. P x}" by (fact frequently_cofinite) lemma Alm_all_def: "Alm_all P \ \ (INFM x. \ P x)" by simp lemma INFM_iff_infinite: "(INFM x. P x) \ infinite {x. P x}" by (fact frequently_cofinite) lemma MOST_iff_cofinite: "(MOST x. P x) \ finite {x. \ P x}" by (fact eventually_cofinite) lemma INFM_EX: "(\\<^sub>\x. P x) \ (\x. P x)" by (fact frequently_ex) lemma ALL_MOST: "\x. P x \ \\<^sub>\x. P x" by (fact always_eventually) lemma INFM_mono: "\\<^sub>\x. P x \ (\x. P x \ Q x) \ \\<^sub>\x. Q x" by (fact frequently_elim1) lemma MOST_mono: "\\<^sub>\x. P x \ (\x. P x \ Q x) \ \\<^sub>\x. Q x" by (fact eventually_mono) lemma INFM_disj_distrib: "(\\<^sub>\x. P x \ Q x) \ (\\<^sub>\x. P x) \ (\\<^sub>\x. Q x)" by (fact frequently_disj_iff) lemma MOST_rev_mp: "\\<^sub>\x. P x \ \\<^sub>\x. P x \ Q x \ \\<^sub>\x. Q x" by (fact eventually_rev_mp) lemma MOST_conj_distrib: "(\\<^sub>\x. P x \ Q x) \ (\\<^sub>\x. P x) \ (\\<^sub>\x. Q x)" by (fact eventually_conj_iff) lemma MOST_conjI: "MOST x. P x \ MOST x. Q x \ MOST x. P x \ Q x" by (fact eventually_conj) lemma INFM_finite_Bex_distrib: "finite A \ (INFM y. \x\A. P x y) \ (\x\A. INFM y. P x y)" by (fact frequently_bex_finite_distrib) lemma MOST_finite_Ball_distrib: "finite A \ (MOST y. \x\A. P x y) \ (\x\A. MOST y. P x y)" by (fact eventually_ball_finite_distrib) lemma INFM_E: "INFM x. P x \ (\x. P x \ thesis) \ thesis" by (fact frequentlyE) lemma MOST_I: "(\x. P x) \ MOST x. P x" by (rule eventuallyI) lemmas MOST_iff_finiteNeg = MOST_iff_cofinite subsection \Enumeration of an Infinite Set\ text \The set's element type must be wellordered (e.g. the natural numbers).\ text \ Could be generalized to \<^prop>\enumerate' S n = (SOME t. t \ s \ finite {s\S. s < t} \ card {s\S. s < t} = n)\. \ primrec (in wellorder) enumerate :: "'a set \ nat \ 'a" where enumerate_0: "enumerate S 0 = (LEAST n. n \ S)" | enumerate_Suc: "enumerate S (Suc n) = enumerate (S - {LEAST n. n \ S}) n" lemma enumerate_Suc': "enumerate S (Suc n) = enumerate (S - {enumerate S 0}) n" by simp lemma enumerate_in_set: "infinite S \ enumerate S n \ S" proof (induct n arbitrary: S) case 0 then show ?case by (fastforce intro: LeastI dest!: infinite_imp_nonempty) next case (Suc n) then show ?case by simp (metis DiffE infinite_remove) qed declare enumerate_0 [simp del] enumerate_Suc [simp del] lemma enumerate_step: "infinite S \ enumerate S n < enumerate S (Suc n)" proof (induction n arbitrary: S) case 0 then have "enumerate S 0 \ enumerate S (Suc 0)" by (simp add: enumerate_0 Least_le enumerate_in_set) moreover have "enumerate (S - {enumerate S 0}) 0 \ S - {enumerate S 0}" by (meson "0.prems" enumerate_in_set infinite_remove) then have "enumerate S 0 \ enumerate (S - {enumerate S 0}) 0" by auto ultimately show ?case by (simp add: enumerate_Suc') next case (Suc n) then show ?case by (simp add: enumerate_Suc') qed lemma enumerate_mono: "m < n \ infinite S \ enumerate S m < enumerate S n" by (induct m n rule: less_Suc_induct) (auto intro: enumerate_step) lemma le_enumerate: assumes S: "infinite S" shows "n \ enumerate S n" using S proof (induct n) case 0 then show ?case by simp next case (Suc n) then have "n \ enumerate S n" by simp also note enumerate_mono[of n "Suc n", OF _ \infinite S\] finally show ?case by simp qed lemma infinite_enumerate: assumes fS: "infinite S" shows "\r::nat\nat. strict_mono r \ (\n. r n \ S)" unfolding strict_mono_def using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by auto lemma enumerate_Suc'': fixes S :: "'a::wellorder set" assumes "infinite S" shows "enumerate S (Suc n) = (LEAST s. s \ S \ enumerate S n < s)" using assms proof (induct n arbitrary: S) case 0 then have "\s \ S. enumerate S 0 \ s" by (auto simp: enumerate.simps intro: Least_le) then show ?case unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"] by (intro arg_cong[where f = Least] ext) auto next case (Suc n S) show ?case using enumerate_mono[OF zero_less_Suc \infinite S\, of n] \infinite S\ apply (subst (1 2) enumerate_Suc') apply (subst Suc) apply (use \infinite S\ in simp) apply (intro arg_cong[where f = Least] ext) apply (auto simp flip: enumerate_Suc') done qed lemma enumerate_Ex: fixes S :: "nat set" assumes S: "infinite S" and s: "s \ S" shows "\n. enumerate S n = s" using s proof (induct s rule: less_induct) case (less s) show ?case proof (cases "\y\S. y < s") case True let ?y = "Max {s'\S. s' < s}" from True have y: "\x. ?y < x \ (\s'\S. s' < s \ s' < x)" by (subst Max_less_iff) auto then have y_in: "?y \ {s'\S. s' < s}" by (intro Max_in) auto with less.hyps[of ?y] obtain n where "enumerate S n = ?y" by auto with S have "enumerate S (Suc n) = s" by (auto simp: y less enumerate_Suc'' intro!: Least_equality) then show ?thesis by auto next case False then have "\t\S. s \ t" by auto with \s \ S\ show ?thesis by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) qed qed lemma inj_enumerate: fixes S :: "'a::wellorder set" assumes S: "infinite S" shows "inj (enumerate S)" unfolding inj_on_def proof clarsimp show "\x y. enumerate S x = enumerate S y \ x = y" by (metis neq_iff enumerate_mono[OF _ \infinite S\]) qed text \To generalise this, we'd need a condition that all initial segments were finite\ lemma bij_enumerate: fixes S :: "nat set" assumes S: "infinite S" shows "bij_betw (enumerate S) UNIV S" proof - have "\s \ S. \i. enumerate S i = s" using enumerate_Ex[OF S] by auto moreover note \infinite S\ inj_enumerate ultimately show ?thesis unfolding bij_betw_def by (auto intro: enumerate_in_set) qed text \A pair of weird and wonderful lemmas from HOL Light.\ lemma finite_transitivity_chain: assumes "finite A" and R: "\x. \ R x x" "\x y z. \R x y; R y z\ \ R x z" and A: "\x. x \ A \ \y. y \ A \ R x y" shows "A = {}" using \finite A\ A proof (induct A) case empty then show ?case by simp next case (insert a A) have False using R(1)[of a] R(2)[of _ a] insert(3,4) by blast thus ?case .. qed corollary Union_maximal_sets: assumes "finite \" shows "\{T \ \. \U\\. \ T \ U} = \\" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by force show "?rhs \ ?lhs" proof (rule Union_subsetI) fix S assume "S \ \" have "{T \ \. S \ T} = {}" if "\ (\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y)" proof - have \
: "\x. x \ \ \ S \ x \ \y. y \ \ \ S \ y \ x \ y" using that by (blast intro: dual_order.trans psubset_imp_subset) show ?thesis proof (rule finite_transitivity_chain [of _ "\T U. S \ T \ T \ U"]) qed (use assms in \auto intro: \
\) qed with \S \ \\ show "\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y" by blast qed qed subsection \Properties of @{term enumerate} on finite sets\ lemma finite_enumerate_in_set: "\finite S; n < card S\ \ enumerate S n \ S" proof (induction n arbitrary: S) case 0 then show ?case by (metis all_not_in_conv card_empty enumerate.simps(1) not_less0 wellorder_Least_lemma(1)) next case (Suc n) show ?case using Suc.prems Suc.IH [of "S - {LEAST n. n \ S}"] apply (simp add: enumerate.simps) by (metis Diff_empty Diff_insert0 Suc_lessD card.remove less_Suc_eq) qed lemma finite_enumerate_step: "\finite S; Suc n < card S\ \ enumerate S n < enumerate S (Suc n)" proof (induction n arbitrary: S) case 0 then have "enumerate S 0 \ enumerate S (Suc 0)" by (simp add: Least_le enumerate.simps(1) finite_enumerate_in_set) moreover have "enumerate (S - {enumerate S 0}) 0 \ S - {enumerate S 0}" by (metis 0 Suc_lessD Suc_less_eq card_Suc_Diff1 enumerate_in_set finite_enumerate_in_set) then have "enumerate S 0 \ enumerate (S - {enumerate S 0}) 0" by auto ultimately show ?case by (simp add: enumerate_Suc') next case (Suc n) then show ?case by (simp add: enumerate_Suc' finite_enumerate_in_set) qed lemma finite_enumerate_mono: "\m < n; finite S; n < card S\ \ enumerate S m < enumerate S n" by (induct m n rule: less_Suc_induct) (auto intro: finite_enumerate_step) lemma finite_le_enumerate: assumes "finite S" "n < card S" shows "n \ enumerate S n" using assms proof (induction n) case 0 then show ?case by simp next case (Suc n) then have "n \ enumerate S n" by simp also note finite_enumerate_mono[of n "Suc n", OF _ \finite S\] finally show ?case using Suc.prems(2) Suc_leI by blast qed lemma finite_enumerate: assumes fS: "finite S" shows "\r::nat\nat. strict_mono_on r {.. (\n S)" unfolding strict_mono_def using finite_enumerate_in_set[OF fS] finite_enumerate_mono[of _ _ S] fS by (metis lessThan_iff strict_mono_on_def) lemma finite_enumerate_Suc'': fixes S :: "'a::wellorder set" assumes "finite S" "Suc n < card S" shows "enumerate S (Suc n) = (LEAST s. s \ S \ enumerate S n < s)" using assms proof (induction n arbitrary: S) case 0 then have "\s \ S. enumerate S 0 \ s" by (auto simp: enumerate.simps intro: Least_le) then show ?case unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"] by (metis Diff_iff dual_order.strict_iff_order singletonD singletonI) next case (Suc n S) then have "Suc n < card (S - {enumerate S 0})" using Suc.prems(2) finite_enumerate_in_set by force then show ?case apply (subst (1 2) enumerate_Suc') apply (simp add: Suc) apply (intro arg_cong[where f = Least] HOL.ext) using finite_enumerate_mono[OF zero_less_Suc \finite S\, of n] Suc.prems by (auto simp flip: enumerate_Suc') qed lemma finite_enumerate_initial_segment: fixes S :: "'a::wellorder set" - assumes "finite S" "s \ S" and n: "n < card (S \ {.. {.. {.. S \ n < s) = (LEAST n. n \ S)" proof (rule Least_equality) have "\t. t \ S \ t < s" by (metis "0" card_gt_0_iff disjoint_iff_not_equal lessThan_iff) then show "(LEAST n. n \ S) \ S \ (LEAST n. n \ S) < s" by (meson LeastI Least_le le_less_trans) qed (simp add: Least_le) then show ?case by (auto simp: enumerate_0) next case (Suc n) then have less_card: "Suc n < card S" by (meson assms(1) card_mono inf_sup_ord(1) leD le_less_linear order.trans) - obtain t where t: "t \ {s \ S. enumerate S n < s}" + obtain T where T: "T \ {s \ S. enumerate S n < s}" by (metis Infinite_Set.enumerate_step enumerate_in_set finite_enumerate_in_set finite_enumerate_step less_card mem_Collect_eq) - have "(LEAST x. x \ S \ x < s \ enumerate S n < x) = (LEAST s. s \ S \ enumerate S n < s)" + have "(LEAST x. x \ S \ x < s \ enumerate S n < x) = (LEAST x. x \ S \ enumerate S n < x)" (is "_ = ?r") proof (intro Least_equality conjI) show "?r \ S" - by (metis (mono_tags, lifting) LeastI mem_Collect_eq t) - show "?r < s" - using not_less_Least [of _ "\t. t \ S \ enumerate S n < t"] Suc assms - by (metis (no_types, lifting) Int_Collect Suc_lessD finite_Int finite_enumerate_in_set finite_enumerate_step lessThan_def linorder_cases) + by (metis (mono_tags, lifting) LeastI mem_Collect_eq T) + have "\ s \ ?r" + using not_less_Least [of _ "\x. x \ S \ enumerate S n < x"] Suc assms + by (metis (mono_tags, lifting) Int_Collect Suc_lessD finite_Int finite_enumerate_in_set finite_enumerate_step lessThan_def less_le_trans) + then show "?r < s" + by auto show "enumerate S n < ?r" - by (metis (no_types, lifting) LeastI mem_Collect_eq t) + by (metis (no_types, lifting) LeastI mem_Collect_eq T) qed (auto simp: Least_le) then show ?case using Suc assms by (simp add: finite_enumerate_Suc'' less_card) qed lemma finite_enumerate_Ex: fixes S :: "'a::wellorder set" assumes S: "finite S" and s: "s \ S" shows "\ny\S. y < s") case True let ?T = "S \ {..finite S\]) from True have y: "\x. Max ?T < x \ (\s'\S. s' < s \ s' < x)" by (subst Max_less_iff) (auto simp: \finite ?T\) then have y_in: "Max ?T \ {s'\S. s' < s}" using Max_in \finite ?T\ by fastforce with less.IH[of "Max ?T" ?T] obtain n where n: "enumerate ?T n = Max ?T" "n < card ?T" using \finite ?T\ by blast then have "Suc n < card S" using TS less_trans_Suc by blast with S n have "enumerate S (Suc n) = s" by (subst finite_enumerate_Suc'') (auto simp: y finite_enumerate_initial_segment less finite_enumerate_Suc'' intro!: Least_equality) then show ?thesis using \Suc n < card S\ by blast next case False then have "\t\S. s \ t" by auto moreover have "0 < card S" using card_0_eq less.prems by blast ultimately show ?thesis using \s \ S\ by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) qed qed lemma finite_bij_enumerate: fixes S :: "'a::wellorder set" assumes S: "finite S" shows "bij_betw (enumerate S) {..n m. \n \ m; n < card S; m < card S\ \ enumerate S n \ enumerate S m" using finite_enumerate_mono[OF _ \finite S\] by (auto simp: neq_iff) then have "inj_on (enumerate S) {..s \ S. \ifinite S\ ultimately show ?thesis unfolding bij_betw_def by (auto intro: finite_enumerate_in_set) qed lemma ex_bij_betw_strict_mono_card: fixes M :: "'a::wellorder set" assumes "finite M" obtains h where "bij_betw h {..