haftmann@35719: (* Title: HOL/Big_Operators.thy wenzelm@12396: Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel avigad@16775: with contributions by Jeremy Avigad wenzelm@12396: *) wenzelm@12396: haftmann@35719: header {* Big operators and finite (non-empty) sets *} haftmann@26041: haftmann@35719: theory Big_Operators haftmann@51489: imports Finite_Set Option Metis haftmann@26041: begin haftmann@26041: haftmann@35816: subsection {* Generic monoid operation over a set *} haftmann@35816: haftmann@35816: no_notation times (infixl "*" 70) haftmann@35816: no_notation Groups.one ("1") haftmann@35816: haftmann@51489: locale comm_monoid_set = comm_monoid haftmann@51489: begin haftmann@35816: haftmann@51738: interpretation comp_fun_commute f haftmann@51738: by default (simp add: fun_eq_iff left_commute) haftmann@51738: haftmann@51738: interpretation comp_fun_commute "f \ g" haftmann@51738: by (rule comp_comp_fun_commute) haftmann@51738: haftmann@51489: definition F :: "('b \ 'a) \ 'b set \ 'a" haftmann@51489: where haftmann@51489: eq_fold: "F g A = Finite_Set.fold (f \ g) 1 A" haftmann@35816: haftmann@35816: lemma infinite [simp]: haftmann@35816: "\ finite A \ F g A = 1" haftmann@51489: by (simp add: eq_fold) haftmann@51489: haftmann@51489: lemma empty [simp]: haftmann@51489: "F g {} = 1" haftmann@51489: by (simp add: eq_fold) haftmann@51489: haftmann@51489: lemma insert [simp]: haftmann@51489: assumes "finite A" and "x \ A" haftmann@51489: shows "F g (insert x A) = g x * F g A" haftmann@51738: using assms by (simp add: eq_fold) haftmann@51489: haftmann@51489: lemma remove: haftmann@51489: assumes "finite A" and "x \ A" haftmann@51489: shows "F g A = g x * F g (A - {x})" haftmann@51489: proof - haftmann@51489: from `x \ A` obtain B where A: "A = insert x B" and "x \ B" haftmann@51489: by (auto dest: mk_disjoint_insert) wenzelm@53374: moreover from `finite A` A have "finite B" by simp haftmann@51489: ultimately show ?thesis by simp haftmann@51489: qed haftmann@51489: haftmann@51489: lemma insert_remove: haftmann@51489: assumes "finite A" haftmann@51489: shows "F g (insert x A) = g x * F g (A - {x})" haftmann@51489: using assms by (cases "x \ A") (simp_all add: remove insert_absorb) haftmann@51489: haftmann@51489: lemma neutral: haftmann@51489: assumes "\x\A. g x = 1" haftmann@51489: shows "F g A = 1" haftmann@51738: using assms by (induct A rule: infinite_finite_induct) simp_all haftmann@35816: haftmann@51489: lemma neutral_const [simp]: haftmann@51489: "F (\_. 1) A = 1" haftmann@51489: by (simp add: neutral) haftmann@51489: haftmann@51489: lemma union_inter: haftmann@51489: assumes "finite A" and "finite B" haftmann@51489: shows "F g (A \ B) * F g (A \ B) = F g A * F g B" haftmann@51489: -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *} haftmann@51489: using assms proof (induct A) haftmann@51489: case empty then show ?case by simp hoelzl@42986: next haftmann@51489: case (insert x A) then show ?case haftmann@51489: by (auto simp add: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute) haftmann@51489: qed haftmann@51489: haftmann@51489: corollary union_inter_neutral: haftmann@51489: assumes "finite A" and "finite B" haftmann@51489: and I0: "\x \ A \ B. g x = 1" haftmann@51489: shows "F g (A \ B) = F g A * F g B" haftmann@51489: using assms by (simp add: union_inter [symmetric] neutral) haftmann@51489: haftmann@51489: corollary union_disjoint: haftmann@51489: assumes "finite A" and "finite B" haftmann@51489: assumes "A \ B = {}" haftmann@51489: shows "F g (A \ B) = F g A * F g B" haftmann@51489: using assms by (simp add: union_inter_neutral) haftmann@51489: haftmann@51489: lemma subset_diff: haftmann@51489: "B \ A \ finite A \ F g A = F g (A - B) * F g B" haftmann@51489: by (metis Diff_partition union_disjoint Diff_disjoint finite_Un inf_commute sup_commute) haftmann@51489: haftmann@51489: lemma reindex: haftmann@51489: assumes "inj_on h A" haftmann@51489: shows "F g (h ` A) = F (g \ h) A" haftmann@51489: proof (cases "finite A") haftmann@51489: case True haftmann@51738: with assms show ?thesis by (simp add: eq_fold fold_image comp_assoc) haftmann@51489: next haftmann@51489: case False with assms have "\ finite (h ` A)" by (blast dest: finite_imageD) haftmann@51489: with False show ?thesis by simp hoelzl@42986: qed hoelzl@42986: haftmann@51489: lemma cong: haftmann@51489: assumes "A = B" haftmann@51489: assumes g_h: "\x. x \ B \ g x = h x" haftmann@51489: shows "F g A = F h B" haftmann@51489: proof (cases "finite A") haftmann@51489: case True haftmann@51489: then have "\C. C \ A \ (\x\C. g x = h x) \ F g C = F h C" haftmann@51489: proof induct haftmann@51489: case empty then show ?case by simp haftmann@51489: next haftmann@51489: case (insert x F) then show ?case apply - haftmann@51489: apply (simp add: subset_insert_iff, clarify) haftmann@51489: apply (subgoal_tac "finite C") haftmann@51489: prefer 2 apply (blast dest: finite_subset [rotated]) haftmann@51489: apply (subgoal_tac "C = insert x (C - {x})") haftmann@51489: prefer 2 apply blast haftmann@51489: apply (erule ssubst) haftmann@51489: apply (simp add: Ball_def del: insert_Diff_single) haftmann@51489: done haftmann@51489: qed haftmann@51489: with `A = B` g_h show ?thesis by simp haftmann@51489: next haftmann@51489: case False haftmann@51489: with `A = B` show ?thesis by simp haftmann@51489: qed nipkow@48849: haftmann@51489: lemma strong_cong [cong]: haftmann@51489: assumes "A = B" "\x. x \ B =simp=> g x = h x" haftmann@51489: shows "F (\x. g x) A = F (\x. h x) B" haftmann@51489: by (rule cong) (insert assms, simp_all add: simp_implies_def) haftmann@51489: haftmann@51489: lemma UNION_disjoint: haftmann@51489: assumes "finite I" and "\i\I. finite (A i)" haftmann@51489: and "\i\I. \j\I. i \ j \ A i \ A j = {}" haftmann@51489: shows "F g (UNION I A) = F (\x. F g (A x)) I" haftmann@51489: apply (insert assms) haftmann@51489: apply (induct rule: finite_induct) haftmann@51489: apply simp haftmann@51489: apply atomize haftmann@51489: apply (subgoal_tac "\i\Fa. x \ i") haftmann@51489: prefer 2 apply blast haftmann@51489: apply (subgoal_tac "A x Int UNION Fa A = {}") haftmann@51489: prefer 2 apply blast haftmann@51489: apply (simp add: union_disjoint) haftmann@51489: done haftmann@51489: haftmann@51489: lemma Union_disjoint: haftmann@51489: assumes "\A\C. finite A" "\A\C. \B\C. A \ B \ A \ B = {}" haftmann@51489: shows "F g (Union C) = F (F g) C" haftmann@51489: proof cases haftmann@51489: assume "finite C" haftmann@51489: from UNION_disjoint [OF this assms] haftmann@51489: show ?thesis haftmann@51489: by (simp add: SUP_def) haftmann@51489: qed (auto dest: finite_UnionD intro: infinite) nipkow@48821: haftmann@51489: lemma distrib: haftmann@51489: "F (\x. g x * h x) A = F g A * F h A" haftmann@51738: using assms by (induct A rule: infinite_finite_induct) (simp_all add: assoc commute left_commute) haftmann@51489: haftmann@51489: lemma Sigma: haftmann@51489: "finite A \ \x\A. finite (B x) \ F (\x. F (g x) (B x)) A = F (split g) (SIGMA x:A. B x)" haftmann@51489: apply (subst Sigma_def) haftmann@51489: apply (subst UNION_disjoint, assumption, simp) haftmann@51489: apply blast haftmann@51489: apply (rule cong) haftmann@51489: apply rule haftmann@51489: apply (simp add: fun_eq_iff) haftmann@51489: apply (subst UNION_disjoint, simp, simp) haftmann@51489: apply blast haftmann@51489: apply (simp add: comp_def) haftmann@51489: done haftmann@51489: haftmann@51489: lemma related: haftmann@51489: assumes Re: "R 1 1" haftmann@51489: and Rop: "\x1 y1 x2 y2. R x1 x2 \ R y1 y2 \ R (x1 * y1) (x2 * y2)" haftmann@51489: and fS: "finite S" and Rfg: "\x\S. R (h x) (g x)" haftmann@51489: shows "R (F h S) (F g S)" haftmann@51489: using fS by (rule finite_subset_induct) (insert assms, auto) nipkow@48849: haftmann@51489: lemma eq_general: haftmann@51489: assumes h: "\y\S'. \!x. x \ S \ h x = y" haftmann@51489: and f12: "\x\S. h x \ S' \ f2 (h x) = f1 x" haftmann@51489: shows "F f1 S = F f2 S'" haftmann@51489: proof- haftmann@51489: from h f12 have hS: "h ` S = S'" by blast haftmann@51489: {fix x y assume H: "x \ S" "y \ S" "h x = h y" haftmann@51489: from f12 h H have "x = y" by auto } haftmann@51489: hence hinj: "inj_on h S" unfolding inj_on_def Ex1_def by blast haftmann@51489: from f12 have th: "\x. x \ S \ (f2 \ h) x = f1 x" by auto haftmann@51489: from hS have "F f2 S' = F f2 (h ` S)" by simp haftmann@51489: also have "\ = F (f2 o h) S" using reindex [OF hinj, of f2] . haftmann@51489: also have "\ = F f1 S " using th cong [of _ _ "f2 o h" f1] haftmann@51489: by blast haftmann@51489: finally show ?thesis .. haftmann@51489: qed nipkow@48849: haftmann@51489: lemma eq_general_reverses: haftmann@51489: assumes kh: "\y. y \ T \ k y \ S \ h (k y) = y" haftmann@51489: and hk: "\x. x \ S \ h x \ T \ k (h x) = x \ g (h x) = j x" haftmann@51489: shows "F j S = F g T" haftmann@51489: (* metis solves it, but not yet available here *) haftmann@51489: apply (rule eq_general [of T S h g j]) haftmann@51489: apply (rule ballI) haftmann@51489: apply (frule kh) haftmann@51489: apply (rule ex1I[]) haftmann@51489: apply blast haftmann@51489: apply clarsimp haftmann@51489: apply (drule hk) apply simp haftmann@51489: apply (rule sym) haftmann@51489: apply (erule conjunct1[OF conjunct2[OF hk]]) haftmann@51489: apply (rule ballI) haftmann@51489: apply (drule hk) haftmann@51489: apply blast haftmann@51489: done haftmann@51489: haftmann@51489: lemma mono_neutral_cong_left: nipkow@48849: assumes "finite T" and "S \ T" and "\i \ T - S. h i = 1" nipkow@48849: and "\x. x \ S \ g x = h x" shows "F g S = F h T" nipkow@48849: proof- nipkow@48849: have eq: "T = S \ (T - S)" using `S \ T` by blast nipkow@48849: have d: "S \ (T - S) = {}" using `S \ T` by blast nipkow@48849: from `finite T` `S \ T` have f: "finite S" "finite (T - S)" nipkow@48849: by (auto intro: finite_subset) nipkow@48849: show ?thesis using assms(4) haftmann@51489: by (simp add: union_disjoint [OF f d, unfolded eq [symmetric]] neutral [OF assms(3)]) nipkow@48849: qed nipkow@48849: haftmann@51489: lemma mono_neutral_cong_right: nipkow@48850: "\ finite T; S \ T; \i \ T - S. g i = 1; \x. x \ S \ g x = h x \ nipkow@48850: \ F g T = F h S" haftmann@51489: by (auto intro!: mono_neutral_cong_left [symmetric]) nipkow@48849: haftmann@51489: lemma mono_neutral_left: nipkow@48849: "\ finite T; S \ T; \i \ T - S. g i = 1 \ \ F g S = F g T" haftmann@51489: by (blast intro: mono_neutral_cong_left) nipkow@48849: haftmann@51489: lemma mono_neutral_right: nipkow@48850: "\ finite T; S \ T; \i \ T - S. g i = 1 \ \ F g T = F g S" haftmann@51489: by (blast intro!: mono_neutral_left [symmetric]) nipkow@48849: haftmann@51489: lemma delta: nipkow@48849: assumes fS: "finite S" haftmann@51489: shows "F (\k. if k = a then b k else 1) S = (if a \ S then b a else 1)" nipkow@48849: proof- nipkow@48849: let ?f = "(\k. if k=a then b k else 1)" nipkow@48849: { assume a: "a \ S" nipkow@48849: hence "\k\S. ?f k = 1" by simp nipkow@48849: hence ?thesis using a by simp } nipkow@48849: moreover nipkow@48849: { assume a: "a \ S" nipkow@48849: let ?A = "S - {a}" nipkow@48849: let ?B = "{a}" nipkow@48849: have eq: "S = ?A \ ?B" using a by blast nipkow@48849: have dj: "?A \ ?B = {}" by simp nipkow@48849: from fS have fAB: "finite ?A" "finite ?B" by auto nipkow@48849: have "F ?f S = F ?f ?A * F ?f ?B" haftmann@51489: using union_disjoint [OF fAB dj, of ?f, unfolded eq [symmetric]] nipkow@48849: by simp haftmann@51489: then have ?thesis using a by simp } nipkow@48849: ultimately show ?thesis by blast nipkow@48849: qed nipkow@48849: haftmann@51489: lemma delta': haftmann@51489: assumes fS: "finite S" haftmann@51489: shows "F (\k. if a = k then b k else 1) S = (if a \ S then b a else 1)" haftmann@51489: using delta [OF fS, of a b, symmetric] by (auto intro: cong) nipkow@48893: hoelzl@42986: lemma If_cases: hoelzl@42986: fixes P :: "'b \ bool" and g h :: "'b \ 'a" hoelzl@42986: assumes fA: "finite A" hoelzl@42986: shows "F (\x. if P x then h x else g x) A = haftmann@51489: F h (A \ {x. P x}) * F g (A \ - {x. P x})" haftmann@51489: proof - hoelzl@42986: have a: "A = A \ {x. P x} \ A \ -{x. P x}" hoelzl@42986: "(A \ {x. P x}) \ (A \ -{x. P x}) = {}" hoelzl@42986: by blast+ hoelzl@42986: from fA hoelzl@42986: have f: "finite (A \ {x. P x})" "finite (A \ -{x. P x})" by auto hoelzl@42986: let ?g = "\x. if P x then h x else g x" haftmann@51489: from union_disjoint [OF f a(2), of ?g] a(1) hoelzl@42986: show ?thesis haftmann@51489: by (subst (1 2) cong) simp_all hoelzl@42986: qed hoelzl@42986: haftmann@51489: lemma cartesian_product: haftmann@51489: "F (\x. F (g x) B) A = F (split g) (A <*> B)" haftmann@51489: apply (rule sym) haftmann@51489: apply (cases "finite A") haftmann@51489: apply (cases "finite B") haftmann@51489: apply (simp add: Sigma) haftmann@51489: apply (cases "A={}", simp) haftmann@51489: apply simp haftmann@51489: apply (auto intro: infinite dest: finite_cartesian_productD2) haftmann@51489: apply (cases "B = {}") apply (auto intro: infinite dest: finite_cartesian_productD1) haftmann@51489: done haftmann@51489: haftmann@35816: end haftmann@35816: haftmann@35816: notation times (infixl "*" 70) haftmann@35816: notation Groups.one ("1") haftmann@35816: haftmann@35816: nipkow@15402: subsection {* Generalized summation over a set *} nipkow@15402: haftmann@51738: context comm_monoid_add haftmann@51738: begin haftmann@51738: haftmann@51738: definition setsum :: "('b \ 'a) \ 'b set \ 'a" haftmann@51489: where haftmann@51489: "setsum = comm_monoid_set.F plus 0" haftmann@26041: haftmann@51738: sublocale setsum!: comm_monoid_set plus 0 haftmann@51489: where haftmann@51546: "comm_monoid_set.F plus 0 = setsum" haftmann@51489: proof - haftmann@51489: show "comm_monoid_set plus 0" .. haftmann@51489: then interpret setsum!: comm_monoid_set plus 0 . haftmann@51546: from setsum_def show "comm_monoid_set.F plus 0 = setsum" by rule haftmann@51489: qed nipkow@15402: wenzelm@19535: abbreviation haftmann@51489: Setsum ("\_" [1000] 999) where haftmann@51489: "\A \ setsum (%x. x) A" wenzelm@19535: haftmann@51738: end haftmann@51738: nipkow@15402: text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is nipkow@15402: written @{text"\x\A. e"}. *} nipkow@15402: nipkow@15402: syntax paulson@17189: "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_. _)" [0, 51, 10] 10) nipkow@15402: syntax (xsymbols) paulson@17189: "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\_\_. _)" [0, 51, 10] 10) nipkow@15402: syntax (HTML output) paulson@17189: "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\_\_. _)" [0, 51, 10] 10) nipkow@15402: nipkow@15402: translations -- {* Beware of argument permutation! *} nipkow@28853: "SUM i:A. b" == "CONST setsum (%i. b) A" nipkow@28853: "\i\A. b" == "CONST setsum (%i. b) A" nipkow@15402: nipkow@15402: text{* Instead of @{term"\x\{x. P}. e"} we introduce the shorter nipkow@15402: @{text"\x|P. e"}. *} nipkow@15402: nipkow@15402: syntax paulson@17189: "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10) nipkow@15402: syntax (xsymbols) paulson@17189: "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) nipkow@15402: syntax (HTML output) paulson@17189: "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) nipkow@15402: nipkow@15402: translations nipkow@28853: "SUM x|P. t" => "CONST setsum (%x. t) {x. P}" nipkow@28853: "\x|P. t" => "CONST setsum (%x. t) {x. P}" nipkow@15402: nipkow@15402: print_translation {* nipkow@15402: let wenzelm@35115: fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] = wenzelm@35115: if x <> y then raise Match wenzelm@35115: else wenzelm@35115: let wenzelm@49660: val x' = Syntax_Trans.mark_bound_body (x, Tx); wenzelm@35115: val t' = subst_bound (x', t); wenzelm@35115: val P' = subst_bound (x', P); wenzelm@49660: in wenzelm@49660: Syntax.const @{syntax_const "_qsetsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t' wenzelm@49660: end wenzelm@35115: | setsum_tr' _ = raise Match; wenzelm@52143: in [(@{const_syntax setsum}, K setsum_tr')] end nipkow@15402: *} nipkow@15402: haftmann@51489: text {* TODO These are candidates for generalization *} nipkow@15402: haftmann@51489: context comm_monoid_add haftmann@51489: begin nipkow@15402: haftmann@51489: lemma setsum_reindex_id: haftmann@35816: "inj_on f B ==> setsum f B = setsum id (f ` B)" haftmann@51489: by (simp add: setsum.reindex) nipkow@15402: haftmann@51489: lemma setsum_reindex_nonzero: chaieb@29674: assumes fS: "finite S" haftmann@51489: and nz: "\x y. x \ S \ y \ S \ x \ y \ f x = f y \ h (f x) = 0" haftmann@51489: shows "setsum h (f ` S) = setsum (h \ f) S" haftmann@51489: using nz proof (induct rule: finite_induct [OF fS]) chaieb@29674: case 1 thus ?case by simp chaieb@29674: next chaieb@29674: case (2 x F) nipkow@48849: { assume fxF: "f x \ f ` F" hence "\y \ F . f y = f x" by auto chaieb@29674: then obtain y where y: "y \ F" "f x = f y" by auto chaieb@29674: from "2.hyps" y have xy: "x \ y" by auto haftmann@51489: from "2.prems" [of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp chaieb@29674: have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto chaieb@29674: also have "\ = setsum (h o f) (insert x F)" haftmann@35816: unfolding setsum.insert[OF `finite F` `x\F`] haftmann@35816: using h0 haftmann@51489: apply (simp cong del: setsum.strong_cong) chaieb@29674: apply (rule "2.hyps"(3)) chaieb@29674: apply (rule_tac y="y" in "2.prems") chaieb@29674: apply simp_all chaieb@29674: done nipkow@48849: finally have ?case . } chaieb@29674: moreover nipkow@48849: { assume fxF: "f x \ f ` F" chaieb@29674: have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)" chaieb@29674: using fxF "2.hyps" by simp chaieb@29674: also have "\ = setsum (h o f) (insert x F)" haftmann@35816: unfolding setsum.insert[OF `finite F` `x\F`] haftmann@51489: apply (simp cong del: setsum.strong_cong) haftmann@35816: apply (rule cong [OF refl [of "op + (h (f x))"]]) chaieb@29674: apply (rule "2.hyps"(3)) chaieb@29674: apply (rule_tac y="y" in "2.prems") chaieb@29674: apply simp_all chaieb@29674: done nipkow@48849: finally have ?case . } chaieb@29674: ultimately show ?case by blast chaieb@29674: qed chaieb@29674: haftmann@51489: lemma setsum_cong2: haftmann@51489: "(\x. x \ A \ f x = g x) \ setsum f A = setsum g A" haftmann@51489: by (auto intro: setsum.cong) nipkow@15554: nipkow@48849: lemma setsum_reindex_cong: nipkow@28853: "[|inj_on f A; B = f ` A; !!a. a:A \ g a = h (f a)|] nipkow@28853: ==> setsum h B = setsum g A" haftmann@51489: by (simp add: setsum.reindex) chaieb@29674: chaieb@30260: lemma setsum_restrict_set: chaieb@30260: assumes fA: "finite A" chaieb@30260: shows "setsum f (A \ B) = setsum (\x. if x \ B then f x else 0) A" chaieb@30260: proof- chaieb@30260: from fA have fab: "finite (A \ B)" by auto chaieb@30260: have aba: "A \ B \ A" by blast chaieb@30260: let ?g = "\x. if x \ A\B then f x else 0" haftmann@51489: from setsum.mono_neutral_left [OF fA aba, of ?g] chaieb@30260: show ?thesis by simp chaieb@30260: qed chaieb@30260: nipkow@15402: lemma setsum_Union_disjoint: hoelzl@44937: assumes "\A\C. finite A" "\A\C. \B\C. A \ B \ A Int B = {}" hoelzl@44937: shows "setsum f (Union C) = setsum (setsum f) C" haftmann@51489: using assms by (fact setsum.Union_disjoint) nipkow@15402: haftmann@51489: lemma setsum_cartesian_product: haftmann@51489: "(\x\A. (\y\B. f x y)) = (\(x,y) \ A <*> B. f x y)" haftmann@51489: by (fact setsum.cartesian_product) nipkow@15402: haftmann@51489: lemma setsum_UNION_zero: nipkow@48893: assumes fS: "finite S" and fSS: "\T \ S. finite T" nipkow@48893: and f0: "\T1 T2 x. T1\S \ T2\S \ T1 \ T2 \ x \ T1 \ x \ T2 \ f x = 0" nipkow@48893: shows "setsum f (\S) = setsum (\T. setsum f T) S" nipkow@48893: using fSS f0 nipkow@48893: proof(induct rule: finite_induct[OF fS]) nipkow@48893: case 1 thus ?case by simp nipkow@48893: next nipkow@48893: case (2 T F) nipkow@48893: then have fTF: "finite T" "\T\F. finite T" "finite F" and TF: "T \ F" nipkow@48893: and H: "setsum f (\ F) = setsum (setsum f) F" by auto nipkow@48893: from fTF have fUF: "finite (\F)" by auto nipkow@48893: from "2.prems" TF fTF nipkow@48893: show ?case haftmann@51489: by (auto simp add: H [symmetric] intro: setsum.union_inter_neutral [OF fTF(1) fUF, of f]) haftmann@51489: qed haftmann@51489: haftmann@51489: text {* Commuting outer and inner summation *} haftmann@51489: haftmann@51489: lemma setsum_commute: haftmann@51489: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" haftmann@51489: proof (simp add: setsum_cartesian_product) haftmann@51489: have "(\(x,y) \ A <*> B. f x y) = haftmann@51489: (\(y,x) \ (%(i, j). (j, i)) ` (A \ B). f x y)" haftmann@51489: (is "?s = _") haftmann@51489: apply (simp add: setsum.reindex [where h = "%(i, j). (j, i)"] swap_inj_on) haftmann@51489: apply (simp add: split_def) haftmann@51489: done haftmann@51489: also have "... = (\(y,x)\B \ A. f x y)" haftmann@51489: (is "_ = ?t") haftmann@51489: apply (simp add: swap_product) haftmann@51489: done haftmann@51489: finally show "?s = ?t" . haftmann@51489: qed haftmann@51489: haftmann@51489: lemma setsum_Plus: haftmann@51489: fixes A :: "'a set" and B :: "'b set" haftmann@51489: assumes fin: "finite A" "finite B" haftmann@51489: shows "setsum f (A <+> B) = setsum (f \ Inl) A + setsum (f \ Inr) B" haftmann@51489: proof - haftmann@51489: have "A <+> B = Inl ` A \ Inr ` B" by auto haftmann@51489: moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)" haftmann@51489: by auto haftmann@51489: moreover have "Inl ` A \ Inr ` B = ({} :: ('a + 'b) set)" by auto haftmann@51489: moreover have "inj_on (Inl :: 'a \ 'a + 'b) A" "inj_on (Inr :: 'b \ 'a + 'b) B" by(auto intro: inj_onI) haftmann@51489: ultimately show ?thesis using fin by(simp add: setsum.union_disjoint setsum.reindex) nipkow@48893: qed nipkow@48893: haftmann@51489: end haftmann@51489: haftmann@51489: text {* TODO These are legacy *} haftmann@51489: haftmann@51489: lemma setsum_empty: haftmann@51489: "setsum f {} = 0" haftmann@51489: by (fact setsum.empty) haftmann@51489: haftmann@51489: lemma setsum_insert: haftmann@51489: "finite F ==> a \ F ==> setsum f (insert a F) = f a + setsum f F" haftmann@51489: by (fact setsum.insert) haftmann@51489: haftmann@51489: lemma setsum_infinite: haftmann@51489: "~ finite A ==> setsum f A = 0" haftmann@51489: by (fact setsum.infinite) haftmann@51489: haftmann@51489: lemma setsum_reindex: haftmann@51489: "inj_on f B \ setsum h (f ` B) = setsum (h \ f) B" haftmann@51489: by (fact setsum.reindex) haftmann@51489: haftmann@51489: lemma setsum_cong: haftmann@51489: "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B" haftmann@51489: by (fact setsum.cong) haftmann@51489: haftmann@51489: lemma strong_setsum_cong: haftmann@51489: "A = B ==> (!!x. x:B =simp=> f x = g x) haftmann@51489: ==> setsum (%x. f x) A = setsum (%x. g x) B" haftmann@51489: by (fact setsum.strong_cong) haftmann@51489: haftmann@51489: lemmas setsum_0 = setsum.neutral_const haftmann@51489: lemmas setsum_0' = setsum.neutral haftmann@51489: haftmann@51489: lemma setsum_Un_Int: "finite A ==> finite B ==> haftmann@51489: setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B" haftmann@51489: -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *} haftmann@51489: by (fact setsum.union_inter) haftmann@51489: haftmann@51489: lemma setsum_Un_disjoint: "finite A ==> finite B haftmann@51489: ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B" haftmann@51489: by (fact setsum.union_disjoint) haftmann@51489: haftmann@51489: lemma setsum_subset_diff: "\ B \ A; finite A \ \ haftmann@51489: setsum f A = setsum f (A - B) + setsum f B" haftmann@51489: by (fact setsum.subset_diff) haftmann@51489: haftmann@51489: lemma setsum_mono_zero_left: haftmann@51489: "\ finite T; S \ T; \i \ T - S. f i = 0 \ \ setsum f S = setsum f T" haftmann@51489: by (fact setsum.mono_neutral_left) haftmann@51489: haftmann@51489: lemmas setsum_mono_zero_right = setsum.mono_neutral_right haftmann@51489: haftmann@51489: lemma setsum_mono_zero_cong_left: haftmann@51489: "\ finite T; S \ T; \i \ T - S. g i = 0; \x. x \ S \ f x = g x \ haftmann@51489: \ setsum f S = setsum g T" haftmann@51489: by (fact setsum.mono_neutral_cong_left) haftmann@51489: haftmann@51489: lemmas setsum_mono_zero_cong_right = setsum.mono_neutral_cong_right haftmann@51489: haftmann@51489: lemma setsum_delta: "finite S \ haftmann@51489: setsum (\k. if k=a then b k else 0) S = (if a \ S then b a else 0)" haftmann@51489: by (fact setsum.delta) haftmann@51489: haftmann@51489: lemma setsum_delta': "finite S \ haftmann@51489: setsum (\k. if a = k then b k else 0) S = (if a\ S then b a else 0)" haftmann@51489: by (fact setsum.delta') haftmann@51489: haftmann@51489: lemma setsum_cases: haftmann@51489: assumes "finite A" haftmann@51489: shows "setsum (\x. if P x then f x else g x) A = haftmann@51489: setsum f (A \ {x. P x}) + setsum g (A \ - {x. P x})" haftmann@51489: using assms by (fact setsum.If_cases) haftmann@51489: haftmann@51489: (*But we can't get rid of finite I. If infinite, although the rhs is 0, haftmann@51489: the lhs need not be, since UNION I A could still be finite.*) haftmann@51489: lemma setsum_UN_disjoint: haftmann@51489: assumes "finite I" and "ALL i:I. finite (A i)" haftmann@51489: and "ALL i:I. ALL j:I. i \ j --> A i Int A j = {}" haftmann@51489: shows "setsum f (UNION I A) = (\i\I. setsum f (A i))" haftmann@51489: using assms by (fact setsum.UNION_disjoint) haftmann@51489: haftmann@51489: (*But we can't get rid of finite A. If infinite, although the lhs is 0, haftmann@51489: the rhs need not be, since SIGMA A B could still be finite.*) haftmann@51489: lemma setsum_Sigma: haftmann@51489: assumes "finite A" and "ALL x:A. finite (B x)" haftmann@51489: shows "(\x\A. (\y\B x. f x y)) = (\(x,y)\(SIGMA x:A. B x). f x y)" haftmann@51489: using assms by (fact setsum.Sigma) haftmann@51489: haftmann@51489: lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)" haftmann@51489: by (fact setsum.distrib) haftmann@51489: haftmann@51489: lemma setsum_Un_zero: haftmann@51489: "\ finite S; finite T; \x \ S\T. f x = 0 \ \ haftmann@51489: setsum f (S \ T) = setsum f S + setsum f T" haftmann@51489: by (fact setsum.union_inter_neutral) haftmann@51489: haftmann@51489: lemma setsum_eq_general_reverses: haftmann@51489: assumes fS: "finite S" and fT: "finite T" haftmann@51489: and kh: "\y. y \ T \ k y \ S \ h (k y) = y" haftmann@51489: and hk: "\x. x \ S \ h x \ T \ k (h x) = x \ g (h x) = f x" haftmann@51489: shows "setsum f S = setsum g T" haftmann@51489: using kh hk by (fact setsum.eq_general_reverses) haftmann@51489: nipkow@15402: nipkow@15402: subsubsection {* Properties in more restricted classes of structures *} nipkow@15402: nipkow@15402: lemma setsum_Un: "finite A ==> finite B ==> nipkow@28853: (setsum f (A Un B) :: 'a :: ab_group_add) = nipkow@28853: setsum f A + setsum f B - setsum f (A Int B)" nipkow@29667: by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps) nipkow@15402: haftmann@49715: lemma setsum_Un2: haftmann@49715: assumes "finite (A \ B)" haftmann@49715: shows "setsum f (A \ B) = setsum f (A - B) + setsum f (B - A) + setsum f (A \ B)" haftmann@49715: proof - haftmann@49715: have "A \ B = A - B \ (B - A) \ A \ B" haftmann@49715: by auto haftmann@49715: with assms show ?thesis by simp (subst setsum_Un_disjoint, auto)+ haftmann@49715: qed haftmann@49715: nipkow@15402: lemma setsum_diff1: "finite A \ nipkow@15402: (setsum f (A - {a}) :: ('a::ab_group_add)) = nipkow@15402: (if a:A then setsum f A - f a else setsum f A)" nipkow@28853: by (erule finite_induct) (auto simp add: insert_Diff_if) nipkow@28853: nipkow@15402: lemma setsum_diff: nipkow@15402: assumes le: "finite A" "B \ A" nipkow@15402: shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))" nipkow@15402: proof - nipkow@15402: from le have finiteB: "finite B" using finite_subset by auto nipkow@15402: show ?thesis using finiteB le wenzelm@21575: proof induct wenzelm@19535: case empty wenzelm@19535: thus ?case by auto wenzelm@19535: next wenzelm@19535: case (insert x F) wenzelm@19535: thus ?case using le finiteB wenzelm@19535: by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb) nipkow@15402: qed wenzelm@19535: qed nipkow@15402: nipkow@15402: lemma setsum_mono: haftmann@35028: assumes le: "\i. i\K \ f (i::'a) \ ((g i)::('b::{comm_monoid_add, ordered_ab_semigroup_add}))" nipkow@15402: shows "(\i\K. f i) \ (\i\K. g i)" nipkow@15402: proof (cases "finite K") nipkow@15402: case True nipkow@15402: thus ?thesis using le wenzelm@19535: proof induct nipkow@15402: case empty nipkow@15402: thus ?case by simp nipkow@15402: next nipkow@15402: case insert nipkow@44890: thus ?case using add_mono by fastforce nipkow@15402: qed nipkow@15402: next haftmann@51489: case False then show ?thesis by simp nipkow@15402: qed nipkow@15402: nipkow@15554: lemma setsum_strict_mono: haftmann@35028: fixes f :: "'a \ 'b::{ordered_cancel_ab_semigroup_add,comm_monoid_add}" wenzelm@19535: assumes "finite A" "A \ {}" wenzelm@19535: and "!!x. x:A \ f x < g x" wenzelm@19535: shows "setsum f A < setsum g A" wenzelm@41550: using assms nipkow@15554: proof (induct rule: finite_ne_induct) nipkow@15554: case singleton thus ?case by simp nipkow@15554: next nipkow@15554: case insert thus ?case by (auto simp: add_strict_mono) nipkow@15554: qed nipkow@15554: nipkow@46699: lemma setsum_strict_mono_ex1: nipkow@46699: fixes f :: "'a \ 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}" nipkow@46699: assumes "finite A" and "ALL x:A. f x \ g x" and "EX a:A. f a < g a" nipkow@46699: shows "setsum f A < setsum g A" nipkow@46699: proof- nipkow@46699: from assms(3) obtain a where a: "a:A" "f a < g a" by blast nipkow@46699: have "setsum f A = setsum f ((A-{a}) \ {a})" nipkow@46699: by(simp add:insert_absorb[OF `a:A`]) nipkow@46699: also have "\ = setsum f (A-{a}) + setsum f {a}" nipkow@46699: using `finite A` by(subst setsum_Un_disjoint) auto nipkow@46699: also have "setsum f (A-{a}) \ setsum g (A-{a})" nipkow@46699: by(rule setsum_mono)(simp add: assms(2)) nipkow@46699: also have "setsum f {a} < setsum g {a}" using a by simp nipkow@46699: also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \ {a})" nipkow@46699: using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto nipkow@46699: also have "\ = setsum g A" by(simp add:insert_absorb[OF `a:A`]) nipkow@46699: finally show ?thesis by (metis add_right_mono add_strict_left_mono) nipkow@46699: qed nipkow@46699: nipkow@15535: lemma setsum_negf: wenzelm@19535: "setsum (%x. - (f x)::'a::ab_group_add) A = - setsum f A" nipkow@15535: proof (cases "finite A") berghofe@22262: case True thus ?thesis by (induct set: finite) auto nipkow@15535: next haftmann@51489: case False thus ?thesis by simp nipkow@15535: qed nipkow@15402: nipkow@15535: lemma setsum_subtractf: wenzelm@19535: "setsum (%x. ((f x)::'a::ab_group_add) - g x) A = wenzelm@19535: setsum f A - setsum g A" nipkow@15535: proof (cases "finite A") nipkow@15535: case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf) nipkow@15535: next haftmann@51489: case False thus ?thesis by simp nipkow@15535: qed nipkow@15402: nipkow@15535: lemma setsum_nonneg: haftmann@35028: assumes nn: "\x\A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \ f x" wenzelm@19535: shows "0 \ setsum f A" nipkow@15535: proof (cases "finite A") nipkow@15535: case True thus ?thesis using nn wenzelm@21575: proof induct wenzelm@19535: case empty then show ?case by simp wenzelm@19535: next wenzelm@19535: case (insert x F) wenzelm@19535: then have "0 + 0 \ f x + setsum f F" by (blast intro: add_mono) wenzelm@19535: with insert show ?case by simp wenzelm@19535: qed nipkow@15535: next haftmann@51489: case False thus ?thesis by simp nipkow@15535: qed nipkow@15402: nipkow@15535: lemma setsum_nonpos: haftmann@35028: assumes np: "\x\A. f x \ (0::'a::{ordered_ab_semigroup_add,comm_monoid_add})" wenzelm@19535: shows "setsum f A \ 0" nipkow@15535: proof (cases "finite A") nipkow@15535: case True thus ?thesis using np wenzelm@21575: proof induct wenzelm@19535: case empty then show ?case by simp wenzelm@19535: next wenzelm@19535: case (insert x F) wenzelm@19535: then have "f x + setsum f F \ 0 + 0" by (blast intro: add_mono) wenzelm@19535: with insert show ?case by simp wenzelm@19535: qed nipkow@15535: next haftmann@51489: case False thus ?thesis by simp nipkow@15535: qed nipkow@15402: hoelzl@36622: lemma setsum_nonneg_leq_bound: hoelzl@36622: fixes f :: "'a \ 'b::{ordered_ab_group_add}" hoelzl@36622: assumes "finite s" "\i. i \ s \ f i \ 0" "(\i \ s. f i) = B" "i \ s" hoelzl@36622: shows "f i \ B" hoelzl@36622: proof - hoelzl@36622: have "0 \ (\ i \ s - {i}. f i)" and "0 \ f i" hoelzl@36622: using assms by (auto intro!: setsum_nonneg) hoelzl@36622: moreover hoelzl@36622: have "(\ i \ s - {i}. f i) + f i = B" hoelzl@36622: using assms by (simp add: setsum_diff1) hoelzl@36622: ultimately show ?thesis by auto hoelzl@36622: qed hoelzl@36622: hoelzl@36622: lemma setsum_nonneg_0: hoelzl@36622: fixes f :: "'a \ 'b::{ordered_ab_group_add}" hoelzl@36622: assumes "finite s" and pos: "\ i. i \ s \ f i \ 0" hoelzl@36622: and "(\ i \ s. f i) = 0" and i: "i \ s" hoelzl@36622: shows "f i = 0" hoelzl@36622: using setsum_nonneg_leq_bound[OF assms] pos[OF i] by auto hoelzl@36622: nipkow@15539: lemma setsum_mono2: haftmann@36303: fixes f :: "'a \ 'b :: ordered_comm_monoid_add" nipkow@15539: assumes fin: "finite B" and sub: "A \ B" and nn: "\b. b \ B-A \ 0 \ f b" nipkow@15539: shows "setsum f A \ setsum f B" nipkow@15539: proof - nipkow@15539: have "setsum f A \ setsum f A + setsum f (B-A)" nipkow@15539: by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def) nipkow@15539: also have "\ = setsum f (A \ (B-A))" using fin finite_subset[OF sub fin] nipkow@15539: by (simp add:setsum_Un_disjoint del:Un_Diff_cancel) nipkow@15539: also have "A \ (B-A) = B" using sub by blast nipkow@15539: finally show ?thesis . nipkow@15539: qed nipkow@15542: avigad@16775: lemma setsum_mono3: "finite B ==> A <= B ==> avigad@16775: ALL x: B - A. haftmann@35028: 0 <= ((f x)::'a::{comm_monoid_add,ordered_ab_semigroup_add}) ==> avigad@16775: setsum f A <= setsum f B" avigad@16775: apply (subgoal_tac "setsum f B = setsum f A + setsum f (B - A)") avigad@16775: apply (erule ssubst) avigad@16775: apply (subgoal_tac "setsum f A + 0 <= setsum f A + setsum f (B - A)") avigad@16775: apply simp avigad@16775: apply (rule add_left_mono) avigad@16775: apply (erule setsum_nonneg) avigad@16775: apply (subst setsum_Un_disjoint [THEN sym]) avigad@16775: apply (erule finite_subset, assumption) avigad@16775: apply (rule finite_subset) avigad@16775: prefer 2 avigad@16775: apply assumption haftmann@32698: apply (auto simp add: sup_absorb2) avigad@16775: done avigad@16775: ballarin@19279: lemma setsum_right_distrib: huffman@22934: fixes f :: "'a => ('b::semiring_0)" nipkow@15402: shows "r * setsum f A = setsum (%n. r * f n) A" nipkow@15402: proof (cases "finite A") nipkow@15402: case True nipkow@15402: thus ?thesis wenzelm@21575: proof induct nipkow@15402: case empty thus ?case by simp nipkow@15402: next webertj@49962: case (insert x A) thus ?case by (simp add: distrib_left) nipkow@15402: qed nipkow@15402: next haftmann@51489: case False thus ?thesis by simp nipkow@15402: qed nipkow@15402: ballarin@17149: lemma setsum_left_distrib: huffman@22934: "setsum f A * (r::'a::semiring_0) = (\n\A. f n * r)" ballarin@17149: proof (cases "finite A") ballarin@17149: case True ballarin@17149: then show ?thesis ballarin@17149: proof induct ballarin@17149: case empty thus ?case by simp ballarin@17149: next webertj@49962: case (insert x A) thus ?case by (simp add: distrib_right) ballarin@17149: qed ballarin@17149: next haftmann@51489: case False thus ?thesis by simp ballarin@17149: qed ballarin@17149: ballarin@17149: lemma setsum_divide_distrib: ballarin@17149: "setsum f A / (r::'a::field) = (\n\A. f n / r)" ballarin@17149: proof (cases "finite A") ballarin@17149: case True ballarin@17149: then show ?thesis ballarin@17149: proof induct ballarin@17149: case empty thus ?case by simp ballarin@17149: next ballarin@17149: case (insert x A) thus ?case by (simp add: add_divide_distrib) ballarin@17149: qed ballarin@17149: next haftmann@51489: case False thus ?thesis by simp ballarin@17149: qed ballarin@17149: nipkow@15535: lemma setsum_abs[iff]: haftmann@35028: fixes f :: "'a => ('b::ordered_ab_group_add_abs)" nipkow@15402: shows "abs (setsum f A) \ setsum (%i. abs(f i)) A" nipkow@15535: proof (cases "finite A") nipkow@15535: case True nipkow@15535: thus ?thesis wenzelm@21575: proof induct nipkow@15535: case empty thus ?case by simp nipkow@15535: next nipkow@15535: case (insert x A) nipkow@15535: thus ?case by (auto intro: abs_triangle_ineq order_trans) nipkow@15535: qed nipkow@15402: next haftmann@51489: case False thus ?thesis by simp nipkow@15402: qed nipkow@15402: nipkow@15535: lemma setsum_abs_ge_zero[iff]: haftmann@35028: fixes f :: "'a => ('b::ordered_ab_group_add_abs)" nipkow@15402: shows "0 \ setsum (%i. abs(f i)) A" nipkow@15535: proof (cases "finite A") nipkow@15535: case True nipkow@15535: thus ?thesis wenzelm@21575: proof induct nipkow@15535: case empty thus ?case by simp nipkow@15535: next huffman@36977: case (insert x A) thus ?case by auto nipkow@15535: qed nipkow@15402: next haftmann@51489: case False thus ?thesis by simp nipkow@15402: qed nipkow@15402: nipkow@15539: lemma abs_setsum_abs[simp]: haftmann@35028: fixes f :: "'a => ('b::ordered_ab_group_add_abs)" nipkow@15539: shows "abs (\a\A. abs(f a)) = (\a\A. abs(f a))" nipkow@15539: proof (cases "finite A") nipkow@15539: case True nipkow@15539: thus ?thesis wenzelm@21575: proof induct nipkow@15539: case empty thus ?case by simp nipkow@15539: next nipkow@15539: case (insert a A) nipkow@15539: hence "\\a\insert a A. \f a\\ = \\f a\ + (\a\A. \f a\)\" by simp nipkow@15539: also have "\ = \\f a\ + \\a\A. \f a\\\" using insert by simp avigad@16775: also have "\ = \f a\ + \\a\A. \f a\\" avigad@16775: by (simp del: abs_of_nonneg) nipkow@15539: also have "\ = (\a\insert a A. \f a\)" using insert by simp nipkow@15539: finally show ?case . nipkow@15539: qed nipkow@15539: next haftmann@51489: case False thus ?thesis by simp nipkow@31080: qed nipkow@31080: haftmann@51489: lemma setsum_diff1'[rule_format]: haftmann@51489: "finite A \ a \ A \ (\ x \ A. f x) = f a + (\ x \ (A - {a}). f x)" haftmann@51489: apply (erule finite_induct[where F=A and P="% A. (a \ A \ (\ x \ A. f x) = f a + (\ x \ (A - {a}). f x))"]) haftmann@51489: apply (auto simp add: insert_Diff_if add_ac) haftmann@51489: done ballarin@17149: haftmann@51489: lemma setsum_diff1_ring: assumes "finite A" "a \ A" haftmann@51489: shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)" haftmann@51489: unfolding setsum_diff1'[OF assms] by auto ballarin@17149: ballarin@19279: lemma setsum_product: huffman@22934: fixes f :: "'a => ('b::semiring_0)" ballarin@19279: shows "setsum f A * setsum g B = (\i\A. \j\B. f i * g j)" ballarin@19279: by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute) ballarin@19279: nipkow@34223: lemma setsum_mult_setsum_if_inj: nipkow@34223: fixes f :: "'a => ('b::semiring_0)" nipkow@34223: shows "inj_on (%(a,b). f a * g b) (A \ B) ==> nipkow@34223: setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}" nipkow@34223: by(auto simp: setsum_product setsum_cartesian_product nipkow@34223: intro!: setsum_reindex_cong[symmetric]) nipkow@34223: haftmann@51489: lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a" haftmann@51489: apply (case_tac "finite A") haftmann@51489: prefer 2 apply simp haftmann@51489: apply (erule rev_mp) haftmann@51489: apply (erule finite_induct, auto) haftmann@51489: done haftmann@51489: haftmann@51489: lemma setsum_eq_0_iff [simp]: haftmann@51489: "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))" haftmann@51489: by (induct set: finite) auto haftmann@51489: haftmann@51489: lemma setsum_eq_Suc0_iff: "finite A \ haftmann@51489: setsum f A = Suc 0 \ (EX a:A. f a = Suc 0 & (ALL b:A. a\b \ f b = 0))" haftmann@51489: apply(erule finite_induct) haftmann@51489: apply (auto simp add:add_is_1) haftmann@51489: done haftmann@51489: haftmann@51489: lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]] haftmann@51489: haftmann@51489: lemma setsum_Un_nat: "finite A ==> finite B ==> haftmann@51489: (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)" haftmann@51489: -- {* For the natural numbers, we have subtraction. *} haftmann@51489: by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps) haftmann@51489: haftmann@51489: lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) = haftmann@51489: (if a:A then setsum f A - f a else setsum f A)" haftmann@51489: apply (case_tac "finite A") haftmann@51489: prefer 2 apply simp haftmann@51489: apply (erule finite_induct) haftmann@51489: apply (auto simp add: insert_Diff_if) haftmann@51489: apply (drule_tac a = a in mk_disjoint_insert, auto) haftmann@51489: done haftmann@51489: haftmann@51489: lemma setsum_diff_nat: haftmann@51489: assumes "finite B" and "B \ A" haftmann@51489: shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)" haftmann@51489: using assms haftmann@51489: proof induct haftmann@51489: show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp haftmann@51489: next haftmann@51489: fix F x assume finF: "finite F" and xnotinF: "x \ F" haftmann@51489: and xFinA: "insert x F \ A" haftmann@51489: and IH: "F \ A \ setsum f (A - F) = setsum f A - setsum f F" haftmann@51489: from xnotinF xFinA have xinAF: "x \ (A - F)" by simp haftmann@51489: from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x" haftmann@51489: by (simp add: setsum_diff1_nat) haftmann@51489: from xFinA have "F \ A" by simp haftmann@51489: with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp haftmann@51489: with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x" haftmann@51489: by simp haftmann@51489: from xnotinF have "A - insert x F = (A - F) - {x}" by auto haftmann@51489: with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x" haftmann@51489: by simp haftmann@51489: from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp haftmann@51489: with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" haftmann@51489: by simp haftmann@51489: thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp haftmann@51489: qed haftmann@51489: haftmann@51600: lemma setsum_comp_morphism: haftmann@51600: assumes "h 0 = 0" and "\x y. h (x + y) = h x + h y" haftmann@51600: shows "setsum (h \ g) A = h (setsum g A)" haftmann@51600: proof (cases "finite A") haftmann@51600: case False then show ?thesis by (simp add: assms) haftmann@51600: next haftmann@51600: case True then show ?thesis by (induct A) (simp_all add: assms) haftmann@51600: qed haftmann@51600: haftmann@51489: haftmann@51489: subsubsection {* Cardinality as special case of @{const setsum} *} haftmann@51489: haftmann@51489: lemma card_eq_setsum: haftmann@51489: "card A = setsum (\x. 1) A" haftmann@51489: proof - haftmann@51489: have "plus \ (\_. Suc 0) = (\_. Suc)" haftmann@51489: by (simp add: fun_eq_iff) haftmann@51489: then have "Finite_Set.fold (plus \ (\_. Suc 0)) = Finite_Set.fold (\_. Suc)" haftmann@51489: by (rule arg_cong) haftmann@51489: then have "Finite_Set.fold (plus \ (\_. Suc 0)) 0 A = Finite_Set.fold (\_. Suc) 0 A" haftmann@51489: by (blast intro: fun_cong) haftmann@51489: then show ?thesis by (simp add: card.eq_fold setsum.eq_fold) haftmann@51489: qed haftmann@51489: haftmann@51489: lemma setsum_constant [simp]: haftmann@51489: "(\x \ A. y) = of_nat (card A) * y" haftmann@35722: apply (cases "finite A") haftmann@35722: apply (erule finite_induct) haftmann@35722: apply (auto simp add: algebra_simps) haftmann@35722: done haftmann@35722: haftmann@35722: lemma setsum_bounded: haftmann@35722: assumes le: "\i. i\A \ f i \ (K::'a::{semiring_1, ordered_ab_semigroup_add})" haftmann@51489: shows "setsum f A \ of_nat (card A) * K" haftmann@35722: proof (cases "finite A") haftmann@35722: case True haftmann@35722: thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp haftmann@35722: next haftmann@51489: case False thus ?thesis by simp haftmann@35722: qed haftmann@35722: haftmann@35722: lemma card_UN_disjoint: haftmann@46629: assumes "finite I" and "\i\I. finite (A i)" haftmann@46629: and "\i\I. \j\I. i \ j \ A i \ A j = {}" haftmann@46629: shows "card (UNION I A) = (\i\I. card(A i))" haftmann@46629: proof - haftmann@46629: have "(\i\I. card (A i)) = (\i\I. \x\A i. 1)" by simp haftmann@46629: with assms show ?thesis by (simp add: card_eq_setsum setsum_UN_disjoint del: setsum_constant) haftmann@46629: qed haftmann@35722: haftmann@35722: lemma card_Union_disjoint: haftmann@35722: "finite C ==> (ALL A:C. finite A) ==> haftmann@35722: (ALL A:C. ALL B:C. A \ B --> A Int B = {}) haftmann@35722: ==> card (Union C) = setsum card C" haftmann@35722: apply (frule card_UN_disjoint [of C id]) hoelzl@44937: apply (simp_all add: SUP_def id_def) haftmann@35722: done haftmann@35722: haftmann@35722: haftmann@35722: subsubsection {* Cardinality of products *} haftmann@35722: haftmann@35722: lemma card_SigmaI [simp]: haftmann@35722: "\ finite A; ALL a:A. finite (B a) \ haftmann@35722: \ card (SIGMA x: A. B x) = (\a\A. card (B a))" haftmann@35722: by(simp add: card_eq_setsum setsum_Sigma del:setsum_constant) haftmann@35722: haftmann@35722: (* haftmann@35722: lemma SigmaI_insert: "y \ A ==> haftmann@35722: (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \ (SIGMA x: A. B x))" haftmann@35722: by auto haftmann@35722: *) haftmann@35722: haftmann@35722: lemma card_cartesian_product: "card (A <*> B) = card(A) * card(B)" haftmann@35722: by (cases "finite A \ finite B") haftmann@35722: (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2) haftmann@35722: haftmann@35722: lemma card_cartesian_product_singleton: "card({x} <*> A) = card(A)" haftmann@35722: by (simp add: card_cartesian_product) haftmann@35722: ballarin@17149: nipkow@15402: subsection {* Generalized product over a set *} nipkow@15402: haftmann@51738: context comm_monoid_mult haftmann@51738: begin haftmann@51738: haftmann@51738: definition setprod :: "('b \ 'a) \ 'b set \ 'a" haftmann@51489: where haftmann@51489: "setprod = comm_monoid_set.F times 1" haftmann@35816: haftmann@51738: sublocale setprod!: comm_monoid_set times 1 haftmann@51489: where haftmann@51546: "comm_monoid_set.F times 1 = setprod" haftmann@51489: proof - haftmann@51489: show "comm_monoid_set times 1" .. haftmann@51489: then interpret setprod!: comm_monoid_set times 1 . haftmann@51546: from setprod_def show "comm_monoid_set.F times 1 = setprod" by rule haftmann@51489: qed nipkow@15402: wenzelm@19535: abbreviation haftmann@51489: Setprod ("\_" [1000] 999) where haftmann@51489: "\A \ setprod (\x. x) A" wenzelm@19535: haftmann@51738: end haftmann@51738: nipkow@15402: syntax paulson@17189: "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3PROD _:_. _)" [0, 51, 10] 10) nipkow@15402: syntax (xsymbols) paulson@17189: "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\_\_. _)" [0, 51, 10] 10) nipkow@15402: syntax (HTML output) paulson@17189: "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\_\_. _)" [0, 51, 10] 10) nipkow@16550: nipkow@16550: translations -- {* Beware of argument permutation! *} nipkow@28853: "PROD i:A. b" == "CONST setprod (%i. b) A" nipkow@28853: "\i\A. b" == "CONST setprod (%i. b) A" nipkow@16550: nipkow@16550: text{* Instead of @{term"\x\{x. P}. e"} we introduce the shorter nipkow@16550: @{text"\x|P. e"}. *} nipkow@16550: nipkow@16550: syntax paulson@17189: "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(3PROD _ |/ _./ _)" [0,0,10] 10) nipkow@16550: syntax (xsymbols) paulson@17189: "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) nipkow@16550: syntax (HTML output) paulson@17189: "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) nipkow@16550: nipkow@15402: translations nipkow@28853: "PROD x|P. t" => "CONST setprod (%x. t) {x. P}" nipkow@28853: "\x|P. t" => "CONST setprod (%x. t) {x. P}" nipkow@16550: haftmann@51489: text {* TODO These are candidates for generalization *} haftmann@51489: haftmann@51489: context comm_monoid_mult haftmann@51489: begin haftmann@51489: haftmann@51489: lemma setprod_reindex_id: haftmann@51489: "inj_on f B ==> setprod f B = setprod id (f ` B)" haftmann@51489: by (auto simp add: setprod.reindex) haftmann@51489: haftmann@51489: lemma setprod_reindex_cong: haftmann@51489: "inj_on f A ==> B = f ` A ==> g = h \ f ==> setprod h B = setprod g A" haftmann@51489: by (frule setprod.reindex, simp) haftmann@51489: haftmann@51489: lemma strong_setprod_reindex_cong: haftmann@51489: assumes i: "inj_on f A" haftmann@51489: and B: "B = f ` A" and eq: "\x. x \ A \ g x = (h \ f) x" haftmann@51489: shows "setprod h B = setprod g A" haftmann@51489: proof- haftmann@51489: have "setprod h B = setprod (h o f) A" haftmann@51489: by (simp add: B setprod.reindex [OF i, of h]) haftmann@51489: then show ?thesis apply simp haftmann@51489: apply (rule setprod.cong) haftmann@51489: apply simp haftmann@51489: by (simp add: eq) haftmann@51489: qed haftmann@51489: haftmann@51489: lemma setprod_Union_disjoint: haftmann@51489: assumes "\A\C. finite A" "\A\C. \B\C. A \ B \ A Int B = {}" haftmann@51489: shows "setprod f (Union C) = setprod (setprod f) C" haftmann@51489: using assms by (fact setprod.Union_disjoint) haftmann@51489: haftmann@51489: text{*Here we can eliminate the finiteness assumptions, by cases.*} haftmann@51489: lemma setprod_cartesian_product: haftmann@51489: "(\x\A. (\y\ B. f x y)) = (\(x,y)\(A <*> B). f x y)" haftmann@51489: by (fact setprod.cartesian_product) haftmann@51489: haftmann@51489: lemma setprod_Un2: haftmann@51489: assumes "finite (A \ B)" haftmann@51489: shows "setprod f (A \ B) = setprod f (A - B) * setprod f (B - A) * setprod f (A \ B)" haftmann@51489: proof - haftmann@51489: have "A \ B = A - B \ (B - A) \ A \ B" haftmann@51489: by auto haftmann@51489: with assms show ?thesis by simp (subst setprod.union_disjoint, auto)+ haftmann@51489: qed haftmann@51489: haftmann@51489: end haftmann@51489: haftmann@51489: text {* TODO These are legacy *} haftmann@51489: haftmann@35816: lemma setprod_empty: "setprod f {} = 1" haftmann@35816: by (fact setprod.empty) nipkow@15402: haftmann@35816: lemma setprod_insert: "[| finite A; a \ A |] ==> nipkow@15402: setprod f (insert a A) = f a * setprod f A" haftmann@35816: by (fact setprod.insert) nipkow@15402: haftmann@35816: lemma setprod_infinite: "~ finite A ==> setprod f A = 1" haftmann@35816: by (fact setprod.infinite) paulson@15409: nipkow@15402: lemma setprod_reindex: haftmann@51489: "inj_on f B ==> setprod h (f ` B) = setprod (h \ f) B" haftmann@51489: by (fact setprod.reindex) nipkow@15402: nipkow@15402: lemma setprod_cong: nipkow@15402: "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B" haftmann@51489: by (fact setprod.cong) nipkow@15402: nipkow@48849: lemma strong_setprod_cong: berghofe@16632: "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B" haftmann@51489: by (fact setprod.strong_cong) nipkow@15402: haftmann@51489: lemma setprod_Un_one: haftmann@51489: "\ finite S; finite T; \x \ S\T. f x = 1 \ haftmann@51489: \ setprod f (S \ T) = setprod f S * setprod f T" haftmann@51489: by (fact setprod.union_inter_neutral) chaieb@29674: haftmann@51489: lemmas setprod_1 = setprod.neutral_const haftmann@51489: lemmas setprod_1' = setprod.neutral nipkow@15402: nipkow@15402: lemma setprod_Un_Int: "finite A ==> finite B nipkow@15402: ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B" haftmann@51489: by (fact setprod.union_inter) nipkow@15402: nipkow@15402: lemma setprod_Un_disjoint: "finite A ==> finite B nipkow@15402: ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B" haftmann@51489: by (fact setprod.union_disjoint) nipkow@48849: nipkow@48849: lemma setprod_subset_diff: "\ B \ A; finite A \ \ nipkow@48849: setprod f A = setprod f (A - B) * setprod f B" haftmann@51489: by (fact setprod.subset_diff) nipkow@15402: nipkow@48849: lemma setprod_mono_one_left: nipkow@48849: "\ finite T; S \ T; \i \ T - S. f i = 1 \ \ setprod f S = setprod f T" haftmann@51489: by (fact setprod.mono_neutral_left) nipkow@30837: haftmann@51489: lemmas setprod_mono_one_right = setprod.mono_neutral_right nipkow@30837: nipkow@48849: lemma setprod_mono_one_cong_left: nipkow@48849: "\ finite T; S \ T; \i \ T - S. g i = 1; \x. x \ S \ f x = g x \ nipkow@48849: \ setprod f S = setprod g T" haftmann@51489: by (fact setprod.mono_neutral_cong_left) nipkow@48849: haftmann@51489: lemmas setprod_mono_one_cong_right = setprod.mono_neutral_cong_right chaieb@29674: nipkow@48849: lemma setprod_delta: "finite S \ nipkow@48849: setprod (\k. if k=a then b k else 1) S = (if a \ S then b a else 1)" haftmann@51489: by (fact setprod.delta) chaieb@29674: nipkow@48849: lemma setprod_delta': "finite S \ nipkow@48849: setprod (\k. if a = k then b k else 1) S = (if a\ S then b a else 1)" haftmann@51489: by (fact setprod.delta') chaieb@29674: nipkow@15402: lemma setprod_UN_disjoint: nipkow@15402: "finite I ==> (ALL i:I. finite (A i)) ==> nipkow@15402: (ALL i:I. ALL j:I. i \ j --> A i Int A j = {}) ==> nipkow@15402: setprod f (UNION I A) = setprod (%i. setprod f (A i)) I" haftmann@51489: by (fact setprod.UNION_disjoint) nipkow@15402: nipkow@15402: lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==> nipkow@16550: (\x\A. (\y\ B x. f x y)) = paulson@17189: (\(x,y)\(SIGMA x:A. B x). f x y)" haftmann@51489: by (fact setprod.Sigma) nipkow@15402: haftmann@51489: lemma setprod_timesf: "setprod (\x. f x * g x) A = setprod f A * setprod g A" haftmann@51489: by (fact setprod.distrib) nipkow@15402: nipkow@15402: nipkow@15402: subsubsection {* Properties in more restricted classes of structures *} nipkow@15402: nipkow@15402: lemma setprod_zero: huffman@23277: "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1) ==> setprod f A = 0" nipkow@28853: apply (induct set: finite, force, clarsimp) nipkow@28853: apply (erule disjE, auto) nipkow@28853: done nipkow@15402: haftmann@51489: lemma setprod_zero_iff[simp]: "finite A ==> haftmann@51489: (setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors})) = haftmann@51489: (EX x: A. f x = 0)" haftmann@51489: by (erule finite_induct, auto simp:no_zero_divisors) haftmann@51489: haftmann@51489: lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \ 0) ==> haftmann@51489: (setprod f (A Un B) :: 'a ::{field}) haftmann@51489: = setprod f A * setprod f B / setprod f (A Int B)" haftmann@51489: by (subst setprod_Un_Int [symmetric], auto) haftmann@51489: nipkow@15402: lemma setprod_nonneg [rule_format]: haftmann@35028: "(ALL x: A. (0::'a::linordered_semidom) \ f x) --> 0 \ setprod f A" huffman@30841: by (cases "finite A", induct set: finite, simp_all add: mult_nonneg_nonneg) huffman@30841: haftmann@35028: lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::linordered_semidom) < f x) nipkow@28853: --> 0 < setprod f A" huffman@30841: by (cases "finite A", induct set: finite, simp_all add: mult_pos_pos) nipkow@15402: nipkow@15402: lemma setprod_diff1: "finite A ==> f a \ 0 ==> nipkow@28853: (setprod f (A - {a}) :: 'a :: {field}) = nipkow@28853: (if a:A then setprod f A / f a else setprod f A)" haftmann@36303: by (erule finite_induct) (auto simp add: insert_Diff_if) nipkow@15402: paulson@31906: lemma setprod_inversef: haftmann@36409: fixes f :: "'b \ 'a::field_inverse_zero" paulson@31906: shows "finite A ==> setprod (inverse \ f) A = inverse (setprod f A)" nipkow@28853: by (erule finite_induct) auto nipkow@15402: nipkow@15402: lemma setprod_dividef: haftmann@36409: fixes f :: "'b \ 'a::field_inverse_zero" wenzelm@31916: shows "finite A nipkow@28853: ==> setprod (%x. f x / g x) A = setprod f A / setprod g A" nipkow@28853: apply (subgoal_tac nipkow@15402: "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \ g) x) A") nipkow@28853: apply (erule ssubst) nipkow@28853: apply (subst divide_inverse) nipkow@28853: apply (subst setprod_timesf) nipkow@28853: apply (subst setprod_inversef, assumption+, rule refl) nipkow@28853: apply (rule setprod_cong, rule refl) nipkow@28853: apply (subst divide_inverse, auto) nipkow@28853: done nipkow@28853: nipkow@29925: lemma setprod_dvd_setprod [rule_format]: nipkow@29925: "(ALL x : A. f x dvd g x) \ setprod f A dvd setprod g A" nipkow@29925: apply (cases "finite A") nipkow@29925: apply (induct set: finite) nipkow@29925: apply (auto simp add: dvd_def) nipkow@29925: apply (rule_tac x = "k * ka" in exI) nipkow@29925: apply (simp add: algebra_simps) nipkow@29925: done nipkow@29925: nipkow@29925: lemma setprod_dvd_setprod_subset: nipkow@29925: "finite B \ A <= B \ setprod f A dvd setprod f B" nipkow@29925: apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)") nipkow@29925: apply (unfold dvd_def, blast) nipkow@29925: apply (subst setprod_Un_disjoint [symmetric]) nipkow@29925: apply (auto elim: finite_subset intro: setprod_cong) nipkow@29925: done nipkow@29925: nipkow@29925: lemma setprod_dvd_setprod_subset2: nipkow@29925: "finite B \ A <= B \ ALL x : A. (f x::'a::comm_semiring_1) dvd g x \ nipkow@29925: setprod f A dvd setprod g B" nipkow@29925: apply (rule dvd_trans) nipkow@29925: apply (rule setprod_dvd_setprod, erule (1) bspec) nipkow@29925: apply (erule (1) setprod_dvd_setprod_subset) nipkow@29925: done nipkow@29925: nipkow@29925: lemma dvd_setprod: "finite A \ i:A \ nipkow@29925: (f i ::'a::comm_semiring_1) dvd setprod f A" nipkow@29925: by (induct set: finite) (auto intro: dvd_mult) nipkow@29925: nipkow@29925: lemma dvd_setsum [rule_format]: "(ALL i : A. d dvd f i) \ nipkow@29925: (d::'a::comm_semiring_1) dvd (SUM x : A. f x)" nipkow@29925: apply (cases "finite A") nipkow@29925: apply (induct set: finite) nipkow@29925: apply auto nipkow@29925: done nipkow@29925: hoelzl@35171: lemma setprod_mono: hoelzl@35171: fixes f :: "'a \ 'b\linordered_semidom" hoelzl@35171: assumes "\i\A. 0 \ f i \ f i \ g i" hoelzl@35171: shows "setprod f A \ setprod g A" hoelzl@35171: proof (cases "finite A") hoelzl@35171: case True hoelzl@35171: hence ?thesis "setprod f A \ 0" using subset_refl[of A] hoelzl@35171: proof (induct A rule: finite_subset_induct) hoelzl@35171: case (insert a F) hoelzl@35171: thus "setprod f (insert a F) \ setprod g (insert a F)" "0 \ setprod f (insert a F)" hoelzl@35171: unfolding setprod_insert[OF insert(1,3)] hoelzl@35171: using assms[rule_format,OF insert(2)] insert hoelzl@35171: by (auto intro: mult_mono mult_nonneg_nonneg) hoelzl@35171: qed auto hoelzl@35171: thus ?thesis by simp hoelzl@35171: qed auto hoelzl@35171: hoelzl@35171: lemma abs_setprod: hoelzl@35171: fixes f :: "'a \ 'b\{linordered_field,abs}" hoelzl@35171: shows "abs (setprod f A) = setprod (\x. abs (f x)) A" hoelzl@35171: proof (cases "finite A") hoelzl@35171: case True thus ?thesis huffman@35216: by induct (auto simp add: field_simps abs_mult) hoelzl@35171: qed auto hoelzl@35171: haftmann@31017: lemma setprod_constant: "finite A ==> (\x\ A. (y::'a::{comm_monoid_mult})) = y^(card A)" nipkow@28853: apply (erule finite_induct) huffman@35216: apply auto nipkow@28853: done nipkow@15402: chaieb@29674: lemma setprod_gen_delta: chaieb@29674: assumes fS: "finite S" haftmann@51489: shows "setprod (\k. if k=a then b k else c) S = (if a \ S then (b a ::'a::comm_monoid_mult) * c^ (card S - 1) else c^ card S)" chaieb@29674: proof- chaieb@29674: let ?f = "(\k. if k=a then b k else c)" chaieb@29674: {assume a: "a \ S" chaieb@29674: hence "\ k\ S. ?f k = c" by simp nipkow@48849: hence ?thesis using a setprod_constant[OF fS, of c] by simp } chaieb@29674: moreover chaieb@29674: {assume a: "a \ S" chaieb@29674: let ?A = "S - {a}" chaieb@29674: let ?B = "{a}" chaieb@29674: have eq: "S = ?A \ ?B" using a by blast chaieb@29674: have dj: "?A \ ?B = {}" by simp chaieb@29674: from fS have fAB: "finite ?A" "finite ?B" by auto chaieb@29674: have fA0:"setprod ?f ?A = setprod (\i. c) ?A" chaieb@29674: apply (rule setprod_cong) by auto chaieb@29674: have cA: "card ?A = card S - 1" using fS a by auto chaieb@29674: have fA1: "setprod ?f ?A = c ^ card ?A" unfolding fA0 apply (rule setprod_constant) using fS by auto chaieb@29674: have "setprod ?f ?A * setprod ?f ?B = setprod ?f S" chaieb@29674: using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]] chaieb@29674: by simp chaieb@29674: then have ?thesis using a cA haftmann@36349: by (simp add: fA1 field_simps cong add: setprod_cong cong del: if_weak_cong)} chaieb@29674: ultimately show ?thesis by blast chaieb@29674: qed chaieb@29674: haftmann@51489: lemma setprod_eq_1_iff [simp]: haftmann@51489: "finite F ==> setprod f F = 1 \ (ALL a:F. f a = (1::nat))" haftmann@51489: by (induct set: finite) auto chaieb@29674: haftmann@51489: lemma setprod_pos_nat: haftmann@51489: "finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0" haftmann@51489: using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric]) haftmann@51489: haftmann@51489: lemma setprod_pos_nat_iff[simp]: haftmann@51489: "finite S ==> (setprod f S > 0) = (ALL x : S. f x > (0::nat))" haftmann@51489: using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric]) haftmann@51489: haftmann@51489: haftmann@51489: subsection {* Generic lattice operations over a set *} haftmann@35816: haftmann@35816: no_notation times (infixl "*" 70) haftmann@35816: no_notation Groups.one ("1") haftmann@35816: haftmann@51489: haftmann@51489: subsubsection {* Without neutral element *} haftmann@51489: haftmann@51489: locale semilattice_set = semilattice haftmann@51489: begin haftmann@51489: haftmann@51738: interpretation comp_fun_idem f haftmann@51738: by default (simp_all add: fun_eq_iff left_commute) haftmann@51738: haftmann@51489: definition F :: "'a set \ 'a" haftmann@51489: where haftmann@51489: eq_fold': "F A = the (Finite_Set.fold (\x y. Some (case y of None \ x | Some z \ f x z)) None A)" haftmann@51489: haftmann@51489: lemma eq_fold: haftmann@51489: assumes "finite A" haftmann@51489: shows "F (insert x A) = Finite_Set.fold f x A" haftmann@51489: proof (rule sym) haftmann@51489: let ?f = "\x y. Some (case y of None \ x | Some z \ f x z)" haftmann@51489: interpret comp_fun_idem "?f" haftmann@51489: by default (simp_all add: fun_eq_iff commute left_commute split: option.split) haftmann@51489: from assms show "Finite_Set.fold f x A = F (insert x A)" haftmann@51489: proof induct haftmann@51489: case empty then show ?case by (simp add: eq_fold') haftmann@51489: next haftmann@51489: case (insert y B) then show ?case by (simp add: insert_commute [of x] eq_fold') haftmann@51489: qed haftmann@51489: qed haftmann@51489: haftmann@51489: lemma singleton [simp]: haftmann@51489: "F {x} = x" haftmann@51489: by (simp add: eq_fold) haftmann@51489: haftmann@51489: lemma insert_not_elem: haftmann@51489: assumes "finite A" and "x \ A" and "A \ {}" haftmann@51489: shows "F (insert x A) = x * F A" haftmann@51489: proof - haftmann@51489: from `A \ {}` obtain b where "b \ A" by blast haftmann@51489: then obtain B where *: "A = insert b B" "b \ B" by (blast dest: mk_disjoint_insert) haftmann@51489: with `finite A` and `x \ A` haftmann@51489: have "finite (insert x B)" and "b \ insert x B" by auto haftmann@51489: then have "F (insert b (insert x B)) = x * F (insert b B)" haftmann@51489: by (simp add: eq_fold) haftmann@51489: then show ?thesis by (simp add: * insert_commute) haftmann@51489: qed haftmann@51489: haftmann@51586: lemma in_idem: haftmann@51489: assumes "finite A" and "x \ A" haftmann@51489: shows "x * F A = F A" haftmann@51489: proof - haftmann@51489: from assms have "A \ {}" by auto haftmann@51489: with `finite A` show ?thesis using `x \ A` haftmann@51489: by (induct A rule: finite_ne_induct) (auto simp add: ac_simps insert_not_elem) haftmann@51489: qed haftmann@51489: haftmann@51489: lemma insert [simp]: haftmann@51489: assumes "finite A" and "A \ {}" haftmann@51489: shows "F (insert x A) = x * F A" haftmann@51586: using assms by (cases "x \ A") (simp_all add: insert_absorb in_idem insert_not_elem) haftmann@51489: haftmann@51489: lemma union: haftmann@51489: assumes "finite A" "A \ {}" and "finite B" "B \ {}" haftmann@51489: shows "F (A \ B) = F A * F B" haftmann@51489: using assms by (induct A rule: finite_ne_induct) (simp_all add: ac_simps) haftmann@51489: haftmann@51489: lemma remove: haftmann@51489: assumes "finite A" and "x \ A" haftmann@51489: shows "F A = (if A - {x} = {} then x else x * F (A - {x}))" haftmann@51489: proof - haftmann@51489: from assms obtain B where "A = insert x B" and "x \ B" by (blast dest: mk_disjoint_insert) haftmann@51489: with assms show ?thesis by simp haftmann@51489: qed haftmann@51489: haftmann@51489: lemma insert_remove: haftmann@51489: assumes "finite A" haftmann@51489: shows "F (insert x A) = (if A - {x} = {} then x else x * F (A - {x}))" haftmann@51489: using assms by (cases "x \ A") (simp_all add: insert_absorb remove) haftmann@51489: haftmann@51489: lemma subset: haftmann@51489: assumes "finite A" "B \ {}" and "B \ A" haftmann@51489: shows "F B * F A = F A" haftmann@51489: proof - haftmann@51489: from assms have "A \ {}" and "finite B" by (auto dest: finite_subset) haftmann@51489: with assms show ?thesis by (simp add: union [symmetric] Un_absorb1) haftmann@51489: qed haftmann@51489: haftmann@51489: lemma closed: haftmann@51489: assumes "finite A" "A \ {}" and elem: "\x y. x * y \ {x, y}" haftmann@51489: shows "F A \ A" haftmann@51489: using `finite A` `A \ {}` proof (induct rule: finite_ne_induct) haftmann@51489: case singleton then show ?case by simp haftmann@51489: next haftmann@51489: case insert with elem show ?case by force haftmann@51489: qed haftmann@51489: haftmann@51489: lemma hom_commute: haftmann@51489: assumes hom: "\x y. h (x * y) = h x * h y" haftmann@51489: and N: "finite N" "N \ {}" haftmann@51489: shows "h (F N) = F (h ` N)" haftmann@51489: using N proof (induct rule: finite_ne_induct) haftmann@51489: case singleton thus ?case by simp haftmann@51489: next haftmann@51489: case (insert n N) haftmann@51489: then have "h (F (insert n N)) = h (n * F N)" by simp haftmann@51489: also have "\ = h n * h (F N)" by (rule hom) haftmann@51489: also have "h (F N) = F (h ` N)" by (rule insert) haftmann@51489: also have "h n * \ = F (insert (h n) (h ` N))" haftmann@51489: using insert by simp haftmann@51489: also have "insert (h n) (h ` N) = h ` insert n N" by simp haftmann@51489: finally show ?case . haftmann@51489: qed haftmann@51489: haftmann@51489: end haftmann@51489: haftmann@51489: locale semilattice_order_set = semilattice_order + semilattice_set haftmann@51489: begin haftmann@51489: haftmann@51489: lemma bounded_iff: haftmann@51489: assumes "finite A" and "A \ {}" haftmann@51489: shows "x \ F A \ (\a\A. x \ a)" haftmann@51489: using assms by (induct rule: finite_ne_induct) (simp_all add: bounded_iff) haftmann@51489: haftmann@51489: lemma boundedI: haftmann@51489: assumes "finite A" haftmann@51489: assumes "A \ {}" haftmann@51489: assumes "\a. a \ A \ x \ a" haftmann@51489: shows "x \ F A" haftmann@51489: using assms by (simp add: bounded_iff) haftmann@51489: haftmann@51489: lemma boundedE: haftmann@51489: assumes "finite A" and "A \ {}" and "x \ F A" haftmann@51489: obtains "\a. a \ A \ x \ a" haftmann@51489: using assms by (simp add: bounded_iff) haftmann@35816: haftmann@51489: lemma coboundedI: haftmann@51489: assumes "finite A" haftmann@51489: and "a \ A" haftmann@51489: shows "F A \ a" haftmann@51489: proof - haftmann@51489: from assms have "A \ {}" by auto haftmann@51489: from `finite A` `A \ {}` `a \ A` show ?thesis haftmann@51489: proof (induct rule: finite_ne_induct) haftmann@51489: case singleton thus ?case by (simp add: refl) haftmann@51489: next haftmann@51489: case (insert x B) haftmann@51489: from insert have "a = x \ a \ B" by simp haftmann@51489: then show ?case using insert by (auto intro: coboundedI2) haftmann@51489: qed haftmann@51489: qed haftmann@51489: haftmann@51489: lemma antimono: haftmann@51489: assumes "A \ B" and "A \ {}" and "finite B" haftmann@51489: shows "F B \ F A" haftmann@51489: proof (cases "A = B") haftmann@51489: case True then show ?thesis by (simp add: refl) haftmann@51489: next haftmann@51489: case False haftmann@51489: have B: "B = A \ (B - A)" using `A \ B` by blast haftmann@51489: then have "F B = F (A \ (B - A))" by simp haftmann@51489: also have "\ = F A * F (B - A)" using False assms by (subst union) (auto intro: finite_subset) haftmann@51489: also have "\ \ F A" by simp haftmann@51489: finally show ?thesis . haftmann@51489: qed haftmann@51489: haftmann@51489: end haftmann@51489: haftmann@51489: haftmann@51489: subsubsection {* With neutral element *} haftmann@51489: haftmann@51489: locale semilattice_neutr_set = semilattice_neutr haftmann@51489: begin haftmann@51489: haftmann@51738: interpretation comp_fun_idem f haftmann@51738: by default (simp_all add: fun_eq_iff left_commute) haftmann@51738: haftmann@51489: definition F :: "'a set \ 'a" haftmann@51489: where haftmann@51489: eq_fold: "F A = Finite_Set.fold f 1 A" haftmann@51489: haftmann@51489: lemma infinite [simp]: haftmann@51489: "\ finite A \ F A = 1" haftmann@51489: by (simp add: eq_fold) haftmann@51489: haftmann@51489: lemma empty [simp]: haftmann@51489: "F {} = 1" haftmann@51489: by (simp add: eq_fold) haftmann@51489: haftmann@51489: lemma insert [simp]: haftmann@51489: assumes "finite A" haftmann@51489: shows "F (insert x A) = x * F A" haftmann@51738: using assms by (simp add: eq_fold) haftmann@51489: haftmann@51586: lemma in_idem: haftmann@51489: assumes "finite A" and "x \ A" haftmann@51489: shows "x * F A = F A" haftmann@51489: proof - haftmann@51489: from assms have "A \ {}" by auto haftmann@51489: with `finite A` show ?thesis using `x \ A` haftmann@51489: by (induct A rule: finite_ne_induct) (auto simp add: ac_simps) haftmann@51489: qed haftmann@51489: haftmann@51489: lemma union: haftmann@51489: assumes "finite A" and "finite B" haftmann@51489: shows "F (A \ B) = F A * F B" haftmann@51489: using assms by (induct A) (simp_all add: ac_simps) haftmann@51489: haftmann@51489: lemma remove: haftmann@51489: assumes "finite A" and "x \ A" haftmann@51489: shows "F A = x * F (A - {x})" haftmann@51489: proof - haftmann@51489: from assms obtain B where "A = insert x B" and "x \ B" by (blast dest: mk_disjoint_insert) haftmann@51489: with assms show ?thesis by simp haftmann@51489: qed haftmann@51489: haftmann@51489: lemma insert_remove: haftmann@51489: assumes "finite A" haftmann@51489: shows "F (insert x A) = x * F (A - {x})" haftmann@51489: using assms by (cases "x \ A") (simp_all add: insert_absorb remove) haftmann@51489: haftmann@51489: lemma subset: haftmann@51489: assumes "finite A" and "B \ A" haftmann@51489: shows "F B * F A = F A" haftmann@51489: proof - haftmann@51489: from assms have "finite B" by (auto dest: finite_subset) haftmann@51489: with assms show ?thesis by (simp add: union [symmetric] Un_absorb1) haftmann@51489: qed haftmann@51489: haftmann@51489: lemma closed: haftmann@51489: assumes "finite A" "A \ {}" and elem: "\x y. x * y \ {x, y}" haftmann@51489: shows "F A \ A" haftmann@51489: using `finite A` `A \ {}` proof (induct rule: finite_ne_induct) haftmann@51489: case singleton then show ?case by simp haftmann@51489: next haftmann@51489: case insert with elem show ?case by force haftmann@51489: qed haftmann@51489: haftmann@51489: end haftmann@51489: haftmann@51489: locale semilattice_order_neutr_set = semilattice_neutr_order + semilattice_neutr_set haftmann@51489: begin haftmann@51489: haftmann@51489: lemma bounded_iff: haftmann@51489: assumes "finite A" haftmann@51489: shows "x \ F A \ (\a\A. x \ a)" haftmann@51489: using assms by (induct A) (simp_all add: bounded_iff) haftmann@51489: haftmann@51489: lemma boundedI: haftmann@51489: assumes "finite A" haftmann@51489: assumes "\a. a \ A \ x \ a" haftmann@51489: shows "x \ F A" haftmann@51489: using assms by (simp add: bounded_iff) haftmann@51489: haftmann@51489: lemma boundedE: haftmann@51489: assumes "finite A" and "x \ F A" haftmann@51489: obtains "\a. a \ A \ x \ a" haftmann@51489: using assms by (simp add: bounded_iff) haftmann@51489: haftmann@51489: lemma coboundedI: haftmann@51489: assumes "finite A" haftmann@51489: and "a \ A" haftmann@51489: shows "F A \ a" haftmann@51489: proof - haftmann@51489: from assms have "A \ {}" by auto haftmann@51489: from `finite A` `A \ {}` `a \ A` show ?thesis haftmann@51489: proof (induct rule: finite_ne_induct) haftmann@51489: case singleton thus ?case by (simp add: refl) haftmann@51489: next haftmann@51489: case (insert x B) haftmann@51489: from insert have "a = x \ a \ B" by simp haftmann@51489: then show ?case using insert by (auto intro: coboundedI2) haftmann@51489: qed haftmann@51489: qed haftmann@51489: haftmann@51489: lemma antimono: haftmann@51489: assumes "A \ B" and "finite B" haftmann@51489: shows "F B \ F A" haftmann@51489: proof (cases "A = B") haftmann@51489: case True then show ?thesis by (simp add: refl) haftmann@51489: next haftmann@51489: case False haftmann@51489: have B: "B = A \ (B - A)" using `A \ B` by blast haftmann@51489: then have "F B = F (A \ (B - A))" by simp haftmann@51489: also have "\ = F A * F (B - A)" using False assms by (subst union) (auto intro: finite_subset) haftmann@51489: also have "\ \ F A" by simp haftmann@51489: finally show ?thesis . haftmann@51489: qed haftmann@51489: haftmann@51489: end haftmann@35816: haftmann@35816: notation times (infixl "*" 70) haftmann@35816: notation Groups.one ("1") haftmann@22917: haftmann@35816: haftmann@51489: subsection {* Lattice operations on finite sets *} haftmann@35816: haftmann@51489: text {* haftmann@51489: For historic reasons, there is the sublocale dependency from @{class distrib_lattice} haftmann@51489: to @{class linorder}. This is badly designed: both should depend on a common abstract haftmann@51489: distributive lattice rather than having this non-subclass dependecy between two haftmann@51489: classes. But for the moment we have to live with it. This forces us to setup haftmann@51489: this sublocale dependency simultaneously with the lattice operations on finite haftmann@51489: sets, to avoid garbage. haftmann@51489: *} haftmann@22917: wenzelm@53174: definition (in semilattice_inf) Inf_fin :: "'a set \ 'a" ("\\<^sub>f\<^sub>i\<^sub>n_" [900] 900) haftmann@51489: where haftmann@51489: "Inf_fin = semilattice_set.F inf" haftmann@26041: wenzelm@53174: definition (in semilattice_sup) Sup_fin :: "'a set \ 'a" ("\\<^sub>f\<^sub>i\<^sub>n_" [900] 900) haftmann@51489: where haftmann@51489: "Sup_fin = semilattice_set.F sup" haftmann@35816: haftmann@51738: context linorder haftmann@51738: begin haftmann@51738: haftmann@51738: definition Min :: "'a set \ 'a" haftmann@51489: where haftmann@51489: "Min = semilattice_set.F min" haftmann@35816: haftmann@51738: definition Max :: "'a set \ 'a" haftmann@51489: where haftmann@51489: "Max = semilattice_set.F max" haftmann@51489: haftmann@51738: sublocale Min!: semilattice_order_set min less_eq less haftmann@51540: + Max!: semilattice_order_set max greater_eq greater haftmann@51540: where haftmann@51540: "semilattice_set.F min = Min" haftmann@51540: and "semilattice_set.F max = Max" haftmann@51540: proof - haftmann@51540: show "semilattice_order_set min less_eq less" by default (auto simp add: min_def) haftmann@52364: then interpret Min!: semilattice_order_set min less_eq less . haftmann@51540: show "semilattice_order_set max greater_eq greater" by default (auto simp add: max_def) haftmann@51540: then interpret Max!: semilattice_order_set max greater_eq greater . haftmann@51540: from Min_def show "semilattice_set.F min = Min" by rule haftmann@51540: from Max_def show "semilattice_set.F max = Max" by rule haftmann@51540: qed haftmann@51540: haftmann@51540: haftmann@51489: text {* An aside: @{const min}/@{const max} on linear orders as special case of @{const inf}/@{const sup} *} haftmann@35816: haftmann@51738: sublocale min_max!: distrib_lattice min less_eq less max haftmann@51489: where haftmann@51489: "semilattice_inf.Inf_fin min = Min" haftmann@51489: and "semilattice_sup.Sup_fin max = Max" haftmann@26041: proof - haftmann@51489: show "class.distrib_lattice min less_eq less max" haftmann@51489: proof haftmann@51489: fix x y z haftmann@51489: show "max x (min y z) = min (max x y) (max x z)" haftmann@51489: by (auto simp add: min_def max_def) haftmann@51489: qed (auto simp add: min_def max_def not_le less_imp_le) haftmann@51489: then interpret min_max!: distrib_lattice min less_eq less max . haftmann@51489: show "semilattice_inf.Inf_fin min = Min" haftmann@51489: by (simp only: min_max.Inf_fin_def Min_def) haftmann@51489: show "semilattice_sup.Sup_fin max = Max" haftmann@51489: by (simp only: min_max.Sup_fin_def Max_def) haftmann@26041: qed haftmann@26041: haftmann@51489: lemmas le_maxI1 = min_max.sup_ge1 haftmann@51489: lemmas le_maxI2 = min_max.sup_ge2 haftmann@51489: haftmann@51489: lemmas min_ac = min_max.inf_assoc min_max.inf_commute haftmann@51540: min.left_commute haftmann@51489: haftmann@51489: lemmas max_ac = min_max.sup_assoc min_max.sup_commute haftmann@51540: max.left_commute haftmann@51489: haftmann@51738: end haftmann@51738: haftmann@51489: haftmann@51489: text {* Lattice operations proper *} haftmann@51489: haftmann@51489: sublocale semilattice_inf < Inf_fin!: semilattice_order_set inf less_eq less haftmann@51489: where haftmann@51546: "semilattice_set.F inf = Inf_fin" haftmann@26757: proof - haftmann@51489: show "semilattice_order_set inf less_eq less" .. haftmann@52364: then interpret Inf_fin!: semilattice_order_set inf less_eq less . haftmann@51546: from Inf_fin_def show "semilattice_set.F inf = Inf_fin" by rule haftmann@26041: qed haftmann@26041: haftmann@51489: sublocale semilattice_sup < Sup_fin!: semilattice_order_set sup greater_eq greater haftmann@51489: where haftmann@51546: "semilattice_set.F sup = Sup_fin" haftmann@51489: proof - haftmann@51489: show "semilattice_order_set sup greater_eq greater" .. haftmann@51489: then interpret Sup_fin!: semilattice_order_set sup greater_eq greater . haftmann@51546: from Sup_fin_def show "semilattice_set.F sup = Sup_fin" by rule haftmann@51489: qed haftmann@35816: haftmann@51489: haftmann@51540: text {* An aside again: @{const Min}/@{const Max} on linear orders as special case of @{const Inf_fin}/@{const Sup_fin} *} haftmann@51540: haftmann@51540: lemma Inf_fin_Min: haftmann@51540: "Inf_fin = (Min :: 'a::{semilattice_inf, linorder} set \ 'a)" haftmann@51540: by (simp add: Inf_fin_def Min_def inf_min) haftmann@51540: haftmann@51540: lemma Sup_fin_Max: haftmann@51540: "Sup_fin = (Max :: 'a::{semilattice_sup, linorder} set \ 'a)" haftmann@51540: by (simp add: Sup_fin_def Max_def sup_max) haftmann@51540: haftmann@51540: haftmann@51540: haftmann@51489: subsection {* Infimum and Supremum over non-empty sets *} haftmann@22917: haftmann@51489: text {* haftmann@51489: After this non-regular bootstrap, things continue canonically. haftmann@51489: *} haftmann@35816: haftmann@35816: context lattice haftmann@35816: begin haftmann@25062: wenzelm@53174: lemma Inf_le_Sup [simp]: "\ finite A; A \ {} \ \ \\<^sub>f\<^sub>i\<^sub>nA \ \\<^sub>f\<^sub>i\<^sub>nA" nipkow@15500: apply(subgoal_tac "EX a. a:A") nipkow@15500: prefer 2 apply blast nipkow@15500: apply(erule exE) haftmann@22388: apply(rule order_trans) haftmann@51489: apply(erule (1) Inf_fin.coboundedI) haftmann@51489: apply(erule (1) Sup_fin.coboundedI) nipkow@15500: done nipkow@15500: haftmann@24342: lemma sup_Inf_absorb [simp]: wenzelm@53174: "finite A \ a \ A \ sup a (\\<^sub>f\<^sub>i\<^sub>nA) = a" nipkow@15512: apply(subst sup_commute) haftmann@51489: apply(simp add: sup_absorb2 Inf_fin.coboundedI) nipkow@15504: done nipkow@15504: haftmann@24342: lemma inf_Sup_absorb [simp]: wenzelm@53174: "finite A \ a \ A \ inf a (\\<^sub>f\<^sub>i\<^sub>nA) = a" haftmann@51489: by (simp add: inf_absorb1 Sup_fin.coboundedI) haftmann@24342: haftmann@24342: end haftmann@24342: haftmann@24342: context distrib_lattice haftmann@24342: begin haftmann@24342: haftmann@24342: lemma sup_Inf1_distrib: haftmann@26041: assumes "finite A" haftmann@26041: and "A \ {}" wenzelm@53174: shows "sup x (\\<^sub>f\<^sub>i\<^sub>nA) = \\<^sub>f\<^sub>i\<^sub>n{sup x a|a. a \ A}" haftmann@51489: using assms by (simp add: image_def Inf_fin.hom_commute [where h="sup x", OF sup_inf_distrib1]) haftmann@51489: (rule arg_cong [where f="Inf_fin"], blast) nipkow@18423: haftmann@24342: lemma sup_Inf2_distrib: haftmann@24342: assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" wenzelm@53174: shows "sup (\\<^sub>f\<^sub>i\<^sub>nA) (\\<^sub>f\<^sub>i\<^sub>nB) = \\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \ A \ b \ B}" haftmann@24342: using A proof (induct rule: finite_ne_induct) haftmann@51489: case singleton then show ?case wenzelm@41550: by (simp add: sup_Inf1_distrib [OF B]) nipkow@15500: next nipkow@15500: case (insert x A) haftmann@25062: have finB: "finite {sup x b |b. b \ B}" haftmann@51489: by (rule finite_surj [where f = "sup x", OF B(1)], auto) haftmann@25062: have finAB: "finite {sup a b |a b. a \ A \ b \ B}" nipkow@15500: proof - haftmann@25062: have "{sup a b |a b. a \ A \ b \ B} = (UN a:A. UN b:B. {sup a b})" nipkow@15500: by blast berghofe@15517: thus ?thesis by(simp add: insert(1) B(1)) nipkow@15500: qed haftmann@25062: have ne: "{sup a b |a b. a \ A \ b \ B} \ {}" using insert B by blast wenzelm@53174: have "sup (\\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\\<^sub>f\<^sub>i\<^sub>nB) = sup (inf x (\\<^sub>f\<^sub>i\<^sub>nA)) (\\<^sub>f\<^sub>i\<^sub>nB)" wenzelm@41550: using insert by simp wenzelm@53174: also have "\ = inf (sup x (\\<^sub>f\<^sub>i\<^sub>nB)) (sup (\\<^sub>f\<^sub>i\<^sub>nA) (\\<^sub>f\<^sub>i\<^sub>nB))" by(rule sup_inf_distrib2) wenzelm@53174: also have "\ = inf (\\<^sub>f\<^sub>i\<^sub>n{sup x b|b. b \ B}) (\\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \ A \ b \ B})" nipkow@15500: using insert by(simp add:sup_Inf1_distrib[OF B]) wenzelm@53174: also have "\ = \\<^sub>f\<^sub>i\<^sub>n({sup x b |b. b \ B} \ {sup a b |a b. a \ A \ b \ B})" wenzelm@53174: (is "_ = \\<^sub>f\<^sub>i\<^sub>n?M") nipkow@15500: using B insert haftmann@51489: by (simp add: Inf_fin.union [OF finB _ finAB ne]) haftmann@25062: also have "?M = {sup a b |a b. a \ insert x A \ b \ B}" nipkow@15500: by blast nipkow@15500: finally show ?case . nipkow@15500: qed nipkow@15500: haftmann@24342: lemma inf_Sup1_distrib: haftmann@26041: assumes "finite A" and "A \ {}" wenzelm@53174: shows "inf x (\\<^sub>f\<^sub>i\<^sub>nA) = \\<^sub>f\<^sub>i\<^sub>n{inf x a|a. a \ A}" haftmann@51489: using assms by (simp add: image_def Sup_fin.hom_commute [where h="inf x", OF inf_sup_distrib1]) haftmann@51489: (rule arg_cong [where f="Sup_fin"], blast) nipkow@18423: haftmann@24342: lemma inf_Sup2_distrib: haftmann@24342: assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" wenzelm@53174: shows "inf (\\<^sub>f\<^sub>i\<^sub>nA) (\\<^sub>f\<^sub>i\<^sub>nB) = \\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \ A \ b \ B}" haftmann@24342: using A proof (induct rule: finite_ne_induct) nipkow@18423: case singleton thus ?case huffman@44921: by(simp add: inf_Sup1_distrib [OF B]) nipkow@18423: next nipkow@18423: case (insert x A) haftmann@25062: have finB: "finite {inf x b |b. b \ B}" haftmann@25062: by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto) haftmann@25062: have finAB: "finite {inf a b |a b. a \ A \ b \ B}" nipkow@18423: proof - haftmann@25062: have "{inf a b |a b. a \ A \ b \ B} = (UN a:A. UN b:B. {inf a b})" nipkow@18423: by blast nipkow@18423: thus ?thesis by(simp add: insert(1) B(1)) nipkow@18423: qed haftmann@25062: have ne: "{inf a b |a b. a \ A \ b \ B} \ {}" using insert B by blast wenzelm@53174: have "inf (\\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\\<^sub>f\<^sub>i\<^sub>nB) = inf (sup x (\\<^sub>f\<^sub>i\<^sub>nA)) (\\<^sub>f\<^sub>i\<^sub>nB)" wenzelm@41550: using insert by simp wenzelm@53174: also have "\ = sup (inf x (\\<^sub>f\<^sub>i\<^sub>nB)) (inf (\\<^sub>f\<^sub>i\<^sub>nA) (\\<^sub>f\<^sub>i\<^sub>nB))" by(rule inf_sup_distrib2) wenzelm@53174: also have "\ = sup (\\<^sub>f\<^sub>i\<^sub>n{inf x b|b. b \ B}) (\\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \ A \ b \ B})" nipkow@18423: using insert by(simp add:inf_Sup1_distrib[OF B]) wenzelm@53174: also have "\ = \\<^sub>f\<^sub>i\<^sub>n({inf x b |b. b \ B} \ {inf a b |a b. a \ A \ b \ B})" wenzelm@53174: (is "_ = \\<^sub>f\<^sub>i\<^sub>n?M") nipkow@18423: using B insert haftmann@51489: by (simp add: Sup_fin.union [OF finB _ finAB ne]) haftmann@25062: also have "?M = {inf a b |a b. a \ insert x A \ b \ B}" nipkow@18423: by blast nipkow@18423: finally show ?case . nipkow@18423: qed nipkow@18423: haftmann@24342: end haftmann@24342: haftmann@35719: context complete_lattice haftmann@35719: begin haftmann@35719: haftmann@35719: lemma Inf_fin_Inf: haftmann@35719: assumes "finite A" and "A \ {}" wenzelm@53174: shows "\\<^sub>f\<^sub>i\<^sub>nA = Inf A" haftmann@35719: proof - haftmann@51489: from assms obtain b B where "A = insert b B" and "finite B" by auto haftmann@51489: then show ?thesis haftmann@51489: by (simp add: Inf_fin.eq_fold inf_Inf_fold_inf inf.commute [of b]) haftmann@35719: qed haftmann@35719: haftmann@35719: lemma Sup_fin_Sup: haftmann@35719: assumes "finite A" and "A \ {}" wenzelm@53174: shows "\\<^sub>f\<^sub>i\<^sub>nA = Sup A" haftmann@35719: proof - haftmann@51489: from assms obtain b B where "A = insert b B" and "finite B" by auto haftmann@51489: then show ?thesis haftmann@51489: by (simp add: Sup_fin.eq_fold sup_Sup_fold_sup sup.commute [of b]) haftmann@35719: qed haftmann@35719: haftmann@35719: end haftmann@35719: haftmann@22917: haftmann@51489: subsection {* Minimum and Maximum over non-empty sets *} haftmann@22917: haftmann@24342: context linorder haftmann@22917: begin haftmann@22917: haftmann@26041: lemma dual_min: haftmann@51489: "ord.min greater_eq = max" wenzelm@46904: by (auto simp add: ord.min_def max_def fun_eq_iff) haftmann@26041: haftmann@51489: lemma dual_max: haftmann@51489: "ord.max greater_eq = min" haftmann@51489: by (auto simp add: ord.max_def min_def fun_eq_iff) haftmann@51489: haftmann@51489: lemma dual_Min: haftmann@51489: "linorder.Min greater_eq = Max" haftmann@26041: proof - haftmann@51489: interpret dual!: linorder greater_eq greater by (fact dual_linorder) haftmann@51489: show ?thesis by (simp add: dual.Min_def dual_min Max_def) haftmann@26041: qed haftmann@26041: haftmann@51489: lemma dual_Max: haftmann@51489: "linorder.Max greater_eq = Min" haftmann@26041: proof - haftmann@51489: interpret dual!: linorder greater_eq greater by (fact dual_linorder) haftmann@51489: show ?thesis by (simp add: dual.Max_def dual_max Min_def) haftmann@26041: qed haftmann@26041: haftmann@51540: lemmas Min_singleton = Min.singleton haftmann@51540: lemmas Max_singleton = Max.singleton haftmann@51540: lemmas Min_insert = Min.insert haftmann@51540: lemmas Max_insert = Max.insert haftmann@51540: lemmas Min_Un = Min.union haftmann@51540: lemmas Max_Un = Max.union haftmann@51540: lemmas hom_Min_commute = Min.hom_commute haftmann@51540: lemmas hom_Max_commute = Max.hom_commute haftmann@26041: paulson@24427: lemma Min_in [simp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "Min A \ A" haftmann@51540: using assms by (auto simp add: min_def Min.closed) nipkow@15392: paulson@24427: lemma Max_in [simp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "Max A \ A" haftmann@51540: using assms by (auto simp add: max_def Max.closed) haftmann@26041: haftmann@26041: lemma Min_le [simp]: haftmann@26757: assumes "finite A" and "x \ A" haftmann@26041: shows "Min A \ x" haftmann@51540: using assms by (fact Min.coboundedI) haftmann@26041: haftmann@26041: lemma Max_ge [simp]: haftmann@26757: assumes "finite A" and "x \ A" haftmann@26041: shows "x \ Max A" haftmann@51540: using assms by (fact Max.coboundedI) haftmann@26041: haftmann@30325: lemma Min_eqI: haftmann@30325: assumes "finite A" haftmann@30325: assumes "\y. y \ A \ y \ x" haftmann@30325: and "x \ A" haftmann@30325: shows "Min A = x" haftmann@30325: proof (rule antisym) haftmann@30325: from `x \ A` have "A \ {}" by auto haftmann@30325: with assms show "Min A \ x" by simp haftmann@30325: next haftmann@30325: from assms show "x \ Min A" by simp haftmann@30325: qed haftmann@30325: haftmann@30325: lemma Max_eqI: haftmann@30325: assumes "finite A" haftmann@30325: assumes "\y. y \ A \ y \ x" haftmann@30325: and "x \ A" haftmann@30325: shows "Max A = x" haftmann@30325: proof (rule antisym) haftmann@30325: from `x \ A` have "A \ {}" by auto haftmann@30325: with assms show "Max A \ x" by simp haftmann@30325: next haftmann@30325: from assms show "x \ Max A" by simp haftmann@30325: qed haftmann@30325: haftmann@51738: context haftmann@51738: fixes A :: "'a set" haftmann@51738: assumes fin_nonempty: "finite A" "A \ {}" haftmann@51738: begin haftmann@51738: haftmann@51489: lemma Min_ge_iff [simp, no_atp]: haftmann@51738: "x \ Min A \ (\a\A. x \ a)" haftmann@51738: using fin_nonempty by (fact Min.bounded_iff) haftmann@51489: haftmann@51489: lemma Max_le_iff [simp, no_atp]: haftmann@51738: "Max A \ x \ (\a\A. a \ x)" haftmann@51738: using fin_nonempty by (fact Max.bounded_iff) haftmann@51489: haftmann@51489: lemma Min_gr_iff [simp, no_atp]: haftmann@51738: "x < Min A \ (\a\A. x < a)" haftmann@51738: using fin_nonempty by (induct rule: finite_ne_induct) simp_all haftmann@51489: haftmann@51489: lemma Max_less_iff [simp, no_atp]: haftmann@51738: "Max A < x \ (\a\A. a < x)" haftmann@51738: using fin_nonempty by (induct rule: finite_ne_induct) simp_all haftmann@51489: haftmann@51489: lemma Min_le_iff [no_atp]: haftmann@51738: "Min A \ x \ (\a\A. a \ x)" haftmann@51738: using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_le_iff_disj) haftmann@51489: haftmann@51489: lemma Max_ge_iff [no_atp]: haftmann@51738: "x \ Max A \ (\a\A. x \ a)" haftmann@51738: using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: le_max_iff_disj) haftmann@51489: haftmann@51489: lemma Min_less_iff [no_atp]: haftmann@51738: "Min A < x \ (\a\A. a < x)" haftmann@51738: using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_less_iff_disj) haftmann@51489: haftmann@51489: lemma Max_gr_iff [no_atp]: haftmann@51738: "x < Max A \ (\a\A. x < a)" haftmann@51738: using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: less_max_iff_disj) haftmann@51738: haftmann@51738: end haftmann@51489: haftmann@26041: lemma Min_antimono: haftmann@26041: assumes "M \ N" and "M \ {}" and "finite N" haftmann@26041: shows "Min N \ Min M" haftmann@51540: using assms by (fact Min.antimono) haftmann@26041: haftmann@26041: lemma Max_mono: haftmann@26041: assumes "M \ N" and "M \ {}" and "finite N" haftmann@26041: shows "Max M \ Max N" haftmann@51540: using assms by (fact Max.antimono) haftmann@51489: haftmann@51489: lemma mono_Min_commute: haftmann@51489: assumes "mono f" haftmann@51489: assumes "finite A" and "A \ {}" haftmann@51489: shows "f (Min A) = Min (f ` A)" haftmann@51489: proof (rule linorder_class.Min_eqI [symmetric]) haftmann@51489: from `finite A` show "finite (f ` A)" by simp haftmann@51489: from assms show "f (Min A) \ f ` A" by simp haftmann@51489: fix x haftmann@51489: assume "x \ f ` A" haftmann@51489: then obtain y where "y \ A" and "x = f y" .. haftmann@51489: with assms have "Min A \ y" by auto haftmann@51489: with `mono f` have "f (Min A) \ f y" by (rule monoE) haftmann@51489: with `x = f y` show "f (Min A) \ x" by simp haftmann@51489: qed haftmann@22917: haftmann@51489: lemma mono_Max_commute: haftmann@51489: assumes "mono f" haftmann@51489: assumes "finite A" and "A \ {}" haftmann@51489: shows "f (Max A) = Max (f ` A)" haftmann@51489: proof (rule linorder_class.Max_eqI [symmetric]) haftmann@51489: from `finite A` show "finite (f ` A)" by simp haftmann@51489: from assms show "f (Max A) \ f ` A" by simp haftmann@51489: fix x haftmann@51489: assume "x \ f ` A" haftmann@51489: then obtain y where "y \ A" and "x = f y" .. haftmann@51489: with assms have "y \ Max A" by auto haftmann@51489: with `mono f` have "f y \ f (Max A)" by (rule monoE) haftmann@51489: with `x = f y` show "x \ f (Max A)" by simp haftmann@51489: qed haftmann@51489: haftmann@51489: lemma finite_linorder_max_induct [consumes 1, case_names empty insert]: haftmann@51489: assumes fin: "finite A" haftmann@51489: and empty: "P {}" haftmann@51489: and insert: "\b A. finite A \ \a\A. a < b \ P A \ P (insert b A)" haftmann@51489: shows "P A" urbanc@36079: using fin empty insert nipkow@32006: proof (induct rule: finite_psubset_induct) urbanc@36079: case (psubset A) urbanc@36079: have IH: "\B. \B < A; P {}; (\A b. \finite A; \a\A. a \ P (insert b A))\ \ P B" by fact urbanc@36079: have fin: "finite A" by fact urbanc@36079: have empty: "P {}" by fact urbanc@36079: have step: "\b A. \finite A; \a\A. a < b; P A\ \ P (insert b A)" by fact krauss@26748: show "P A" haftmann@26757: proof (cases "A = {}") urbanc@36079: assume "A = {}" urbanc@36079: then show "P A" using `P {}` by simp krauss@26748: next urbanc@36079: let ?B = "A - {Max A}" urbanc@36079: let ?A = "insert (Max A) ?B" urbanc@36079: have "finite ?B" using `finite A` by simp krauss@26748: assume "A \ {}" krauss@26748: with `finite A` have "Max A : A" by auto urbanc@36079: then have A: "?A = A" using insert_Diff_single insert_absorb by auto haftmann@51489: then have "P ?B" using `P {}` step IH [of ?B] by blast urbanc@36079: moreover nipkow@44890: have "\a\?B. a < Max A" using Max_ge [OF `finite A`] by fastforce haftmann@51489: ultimately show "P A" using A insert_Diff_single step [OF `finite ?B`] by fastforce krauss@26748: qed krauss@26748: qed krauss@26748: haftmann@51489: lemma finite_linorder_min_induct [consumes 1, case_names empty insert]: haftmann@51489: "\finite A; P {}; \b A. \finite A; \a\A. b < a; P A\ \ P (insert b A)\ \ P A" haftmann@51489: by (rule linorder.finite_linorder_max_induct [OF dual_linorder]) nipkow@32006: haftmann@52379: lemma Least_Min: haftmann@52379: assumes "finite {a. P a}" and "\a. P a" haftmann@52379: shows "(LEAST a. P a) = Min {a. P a}" haftmann@52379: proof - haftmann@52379: { fix A :: "'a set" haftmann@52379: assume A: "finite A" "A \ {}" haftmann@52379: have "(LEAST a. a \ A) = Min A" haftmann@52379: using A proof (induct A rule: finite_ne_induct) haftmann@52379: case singleton show ?case by (rule Least_equality) simp_all haftmann@52379: next haftmann@52379: case (insert a A) haftmann@52379: have "(LEAST b. b = a \ b \ A) = min a (LEAST a. a \ A)" haftmann@52379: by (auto intro!: Least_equality simp add: min_def not_le Min_le_iff insert.hyps dest!: less_imp_le) haftmann@52379: with insert show ?case by simp haftmann@52379: qed haftmann@52379: } from this [of "{a. P a}"] assms show ?thesis by simp haftmann@52379: qed haftmann@52379: haftmann@22917: end haftmann@22917: haftmann@35028: context linordered_ab_semigroup_add haftmann@22917: begin haftmann@22917: haftmann@22917: lemma add_Min_commute: haftmann@22917: fixes k haftmann@25062: assumes "finite N" and "N \ {}" haftmann@25062: shows "k + Min N = Min {k + m | m. m \ N}" haftmann@25062: proof - haftmann@25062: have "\x y. k + min x y = min (k + x) (k + y)" haftmann@25062: by (simp add: min_def not_le) haftmann@25062: (blast intro: antisym less_imp_le add_left_mono) haftmann@25062: with assms show ?thesis haftmann@25062: using hom_Min_commute [of "plus k" N] haftmann@25062: by simp (blast intro: arg_cong [where f = Min]) haftmann@25062: qed haftmann@22917: haftmann@22917: lemma add_Max_commute: haftmann@22917: fixes k haftmann@25062: assumes "finite N" and "N \ {}" haftmann@25062: shows "k + Max N = Max {k + m | m. m \ N}" haftmann@25062: proof - haftmann@25062: have "\x y. k + max x y = max (k + x) (k + y)" haftmann@25062: by (simp add: max_def not_le) haftmann@25062: (blast intro: antisym less_imp_le add_left_mono) haftmann@25062: with assms show ?thesis haftmann@25062: using hom_Max_commute [of "plus k" N] haftmann@25062: by simp (blast intro: arg_cong [where f = Max]) haftmann@25062: qed haftmann@22917: haftmann@22917: end haftmann@22917: haftmann@35034: context linordered_ab_group_add haftmann@35034: begin haftmann@35034: haftmann@35034: lemma minus_Max_eq_Min [simp]: haftmann@51489: "finite S \ S \ {} \ - Max S = Min (uminus ` S)" haftmann@35034: by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min) haftmann@35034: haftmann@35034: lemma minus_Min_eq_Max [simp]: haftmann@51489: "finite S \ S \ {} \ - Min S = Max (uminus ` S)" haftmann@35034: by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max) haftmann@35034: haftmann@35034: end haftmann@35034: haftmann@51540: context complete_linorder haftmann@51540: begin haftmann@51540: haftmann@51540: lemma Min_Inf: haftmann@51540: assumes "finite A" and "A \ {}" haftmann@51540: shows "Min A = Inf A" haftmann@51540: proof - haftmann@51540: from assms obtain b B where "A = insert b B" and "finite B" by auto haftmann@51540: then show ?thesis haftmann@51540: by (simp add: Min.eq_fold complete_linorder_inf_min [symmetric] inf_Inf_fold_inf inf.commute [of b]) haftmann@51540: qed haftmann@51540: haftmann@51540: lemma Max_Sup: haftmann@51540: assumes "finite A" and "A \ {}" haftmann@51540: shows "Max A = Sup A" haftmann@51540: proof - haftmann@51540: from assms obtain b B where "A = insert b B" and "finite B" by auto haftmann@51540: then show ?thesis haftmann@51540: by (simp add: Max.eq_fold complete_linorder_sup_max [symmetric] sup_Sup_fold_sup sup.commute [of b]) haftmann@51540: qed haftmann@51540: haftmann@25571: end haftmann@51263: haftmann@51540: end