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@35722: imports Plain 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@35816: locale comm_monoid_big = comm_monoid + haftmann@35816: fixes F :: "('b \ 'a) \ 'b set \ 'a" haftmann@35816: assumes F_eq: "F g A = (if finite A then fold_image (op *) g 1 A else 1)" haftmann@35816: haftmann@35816: sublocale comm_monoid_big < folding_image proof haftmann@35816: qed (simp add: F_eq) haftmann@35816: haftmann@35816: context comm_monoid_big haftmann@35816: begin haftmann@35816: haftmann@35816: lemma infinite [simp]: haftmann@35816: "\ finite A \ F g A = 1" haftmann@35816: by (simp add: F_eq) haftmann@35816: hoelzl@42986: lemma F_cong: hoelzl@42986: assumes "A = B" "\x. x \ B \ h x = g x" hoelzl@42986: shows "F h A = F g B" hoelzl@42986: proof cases hoelzl@42986: assume "finite A" hoelzl@42986: with assms show ?thesis unfolding `A = B` by (simp cong: cong) hoelzl@42986: next hoelzl@42986: assume "\ finite A" hoelzl@42986: then show ?thesis unfolding `A = B` by simp hoelzl@42986: qed hoelzl@42986: nipkow@48849: lemma strong_F_cong [cong]: nipkow@48849: "\ A = B; !!x. x:B =simp=> g x = h x \ nipkow@48849: \ F (%x. g x) A = F (%x. h x) B" nipkow@48849: by (rule F_cong) (simp_all add: simp_implies_def) nipkow@48849: nipkow@48821: lemma F_neutral[simp]: "F (%i. 1) A = 1" nipkow@48821: by (cases "finite A") (simp_all add: neutral) nipkow@48821: nipkow@48821: lemma F_neutral': "ALL a:A. g a = 1 \ F g A = 1" nipkow@48849: by simp nipkow@48849: nipkow@48849: lemma F_subset_diff: "\ B \ A; finite A \ \ F g A = F g (A - B) * F g B" nipkow@48849: by (metis Diff_partition union_disjoint Diff_disjoint finite_Un inf_commute sup_commute) nipkow@48849: nipkow@48849: lemma F_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) nipkow@48849: by (simp add: union_disjoint[OF f d, unfolded eq[symmetric]] F_neutral'[OF assms(3)]) nipkow@48849: qed nipkow@48849: nipkow@48850: lemma F_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" nipkow@48850: by(auto intro!: F_mono_neutral_cong_left[symmetric]) nipkow@48849: nipkow@48849: lemma F_mono_neutral_left: nipkow@48849: "\ finite T; S \ T; \i \ T - S. g i = 1 \ \ F g S = F g T" nipkow@48849: by(blast intro: F_mono_neutral_cong_left) nipkow@48849: nipkow@48850: lemma F_mono_neutral_right: nipkow@48850: "\ finite T; S \ T; \i \ T - S. g i = 1 \ \ F g T = F g S" nipkow@48850: by(blast intro!: F_mono_neutral_left[symmetric]) nipkow@48849: nipkow@48849: lemma F_delta: nipkow@48849: assumes fS: "finite S" nipkow@48849: 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" nipkow@48849: using union_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]] nipkow@48849: by simp nipkow@48849: then have ?thesis using a by simp } nipkow@48849: ultimately show ?thesis by blast nipkow@48849: qed nipkow@48849: nipkow@48849: lemma F_delta': nipkow@48849: assumes fS: "finite S" shows nipkow@48849: "F (\k. if a = k then b k else 1) S = (if a \ S then b a else 1)" nipkow@48849: using F_delta[OF fS, of a b, symmetric] by (auto intro: F_cong) nipkow@48821: nipkow@48861: lemma F_fun_f: "F (%x. g x * h x) A = (F g A * F h A)" nipkow@48861: by (cases "finite A") (simp_all add: distrib) nipkow@48861: nipkow@48893: nipkow@48893: text {* for ad-hoc proofs for @{const fold_image} *} nipkow@48893: lemma comm_monoid_mult: "class.comm_monoid_mult (op *) 1" nipkow@48893: proof qed (auto intro: assoc commute) nipkow@48893: nipkow@48893: lemma F_Un_neutral: nipkow@48893: assumes fS: "finite S" and fT: "finite T" nipkow@48893: and I1: "\x \ S\T. g x = 1" nipkow@48893: shows "F g (S \ T) = F g S * F g T" nipkow@48893: proof - nipkow@48893: interpret comm_monoid_mult "op *" 1 by (fact comm_monoid_mult) nipkow@48893: show ?thesis nipkow@48893: using fS fT nipkow@48893: apply (simp add: F_eq) nipkow@48893: apply (rule fold_image_Un_one) nipkow@48893: using I1 by auto nipkow@48893: qed 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 = hoelzl@42986: F h (A \ {x. P x}) * F g (A \ - {x. P x})" hoelzl@42986: 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" hoelzl@42986: from union_disjoint[OF f a(2), of ?g] a(1) hoelzl@42986: show ?thesis hoelzl@42986: by (subst (1 2) F_cong) simp_all hoelzl@42986: qed hoelzl@42986: haftmann@35816: end haftmann@35816: haftmann@35816: text {* for ad-hoc proofs for @{const fold_image} *} haftmann@35816: haftmann@35816: lemma (in comm_monoid_add) comm_monoid_mult: haftmann@36635: "class.comm_monoid_mult (op +) 0" haftmann@35816: proof qed (auto intro: add_assoc add_commute) 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@35816: definition (in comm_monoid_add) setsum :: "('b \ 'a) => 'b set => 'a" where haftmann@35816: "setsum f A = (if finite A then fold_image (op +) f 0 A else 0)" haftmann@26041: haftmann@35816: sublocale comm_monoid_add < setsum!: comm_monoid_big "op +" 0 setsum proof haftmann@35816: qed (fact setsum_def) nipkow@15402: wenzelm@19535: abbreviation wenzelm@21404: Setsum ("\_" [1000] 999) where wenzelm@19535: "\A == setsum (%x. x) A" wenzelm@19535: 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@42284: val x' = Syntax_Trans.mark_bound x; wenzelm@35115: val t' = subst_bound (x', t); wenzelm@35115: val P' = subst_bound (x', P); wenzelm@42284: in Syntax.const @{syntax_const "_qsetsum"} $ Syntax_Trans.mark_bound x $ P' $ t' end wenzelm@35115: | setsum_tr' _ = raise Match; wenzelm@35115: in [(@{const_syntax setsum}, setsum_tr')] end nipkow@15402: *} nipkow@15402: haftmann@35816: lemma setsum_empty: haftmann@35816: "setsum f {} = 0" haftmann@35816: by (fact setsum.empty) nipkow@15402: haftmann@35816: lemma setsum_insert: nipkow@28853: "finite F ==> a \ F ==> setsum f (insert a F) = f a + setsum f F" haftmann@35816: by (fact setsum.insert) haftmann@35816: haftmann@35816: lemma setsum_infinite: haftmann@35816: "~ finite A ==> setsum f A = 0" haftmann@35816: by (fact setsum.infinite) nipkow@15402: haftmann@35816: lemma (in comm_monoid_add) setsum_reindex: haftmann@35816: assumes "inj_on f B" shows "setsum h (f ` B) = setsum (h \ f) B" haftmann@35816: proof - haftmann@35816: interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult) nipkow@48849: from assms show ?thesis by (auto simp add: setsum_def fold_image_reindex o_def dest!:finite_imageD) haftmann@35816: qed paulson@15409: nipkow@48849: lemma setsum_reindex_id: haftmann@35816: "inj_on f B ==> setsum f B = setsum id (f ` B)" nipkow@48849: by (simp add: setsum_reindex) nipkow@15402: nipkow@48849: lemma setsum_reindex_nonzero: chaieb@29674: assumes fS: "finite S" chaieb@29674: and nz: "\ x y. x \ S \ y \ S \ x \ y \ f x = f y \ h (f x) = 0" chaieb@29674: shows "setsum h (f ` S) = setsum (h o f) S" chaieb@29674: using nz chaieb@29674: 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 chaieb@29674: chaieb@29674: 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 nipkow@48849: apply (simp cong del:setsum.strong_F_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`] nipkow@48849: apply (simp cong del:setsum.strong_F_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: nipkow@48849: lemma setsum_cong: nipkow@15402: "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B" nipkow@48849: by (fact setsum.F_cong) nipkow@15402: nipkow@48849: lemma strong_setsum_cong: nipkow@16733: "A = B ==> (!!x. x:B =simp=> f x = g x) nipkow@16733: ==> setsum (%x. f x) A = setsum (%x. g x) B" nipkow@48849: by (fact setsum.strong_F_cong) berghofe@16632: nipkow@48849: lemma setsum_cong2: "\\x. x \ A \ f x = g x\ \ setsum f A = setsum g A" nipkow@48849: 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" nipkow@48849: by (simp add: setsum_reindex) nipkow@15402: nipkow@48821: lemmas setsum_0 = setsum.F_neutral nipkow@48821: lemmas setsum_0' = setsum.F_neutral' nipkow@15402: nipkow@48849: lemma setsum_Un_Int: "finite A ==> finite B ==> nipkow@15402: setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B" nipkow@15402: -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *} nipkow@48849: by (fact setsum.union_inter) nipkow@15402: nipkow@48849: lemma setsum_Un_disjoint: "finite A ==> finite B nipkow@15402: ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B" nipkow@48849: by (fact setsum.union_disjoint) nipkow@48849: nipkow@48849: lemma setsum_subset_diff: "\ B \ A; finite A \ \ nipkow@48849: setsum f A = setsum f (A - B) + setsum f B" nipkow@48849: by(fact setsum.F_subset_diff) nipkow@15402: chaieb@29674: lemma setsum_mono_zero_left: nipkow@48849: "\ finite T; S \ T; \i \ T - S. f i = 0 \ \ setsum f S = setsum f T" nipkow@48849: by(fact setsum.F_mono_neutral_left) chaieb@29674: nipkow@48849: lemmas setsum_mono_zero_right = setsum.F_mono_neutral_right chaieb@29674: chaieb@29674: lemma setsum_mono_zero_cong_left: nipkow@48849: "\ finite T; S \ T; \i \ T - S. g i = 0; \x. x \ S \ f x = g x \ nipkow@48849: \ setsum f S = setsum g T" nipkow@48849: by(fact setsum.F_mono_neutral_cong_left) chaieb@29674: nipkow@48849: lemmas setsum_mono_zero_cong_right = setsum.F_mono_neutral_cong_right chaieb@29674: nipkow@48849: lemma setsum_delta: "finite S \ nipkow@48849: setsum (\k. if k=a then b k else 0) S = (if a \ S then b a else 0)" nipkow@48849: by(fact setsum.F_delta) nipkow@48849: nipkow@48849: lemma setsum_delta': "finite S \ nipkow@48849: setsum (\k. if a = k then b k else 0) S = (if a\ S then b a else 0)" nipkow@48849: by(fact setsum.F_delta') 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" chaieb@30260: from setsum_mono_zero_left[OF fA aba, of ?g] chaieb@30260: show ?thesis by simp chaieb@30260: qed chaieb@30260: chaieb@30260: lemma setsum_cases: chaieb@30260: assumes fA: "finite A" hoelzl@35577: shows "setsum (\x. if P x then f x else g x) A = hoelzl@35577: setsum f (A \ {x. P x}) + setsum g (A \ - {x. P x})" hoelzl@42986: using setsum.If_cases[OF fA] . chaieb@29674: paulson@15409: (*But we can't get rid of finite I. If infinite, although the rhs is 0, paulson@15409: the lhs need not be, since UNION I A could still be finite.*) haftmann@35816: lemma (in comm_monoid_add) setsum_UN_disjoint: haftmann@35816: assumes "finite I" and "ALL i:I. finite (A i)" haftmann@35816: and "ALL i:I. ALL j:I. i \ j --> A i Int A j = {}" haftmann@35816: shows "setsum f (UNION I A) = (\i\I. setsum f (A i))" haftmann@35816: proof - haftmann@35816: interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult) wenzelm@41550: from assms show ?thesis by (simp add: setsum_def fold_image_UN_disjoint) haftmann@35816: qed nipkow@15402: paulson@15409: text{*No need to assume that @{term C} is finite. If infinite, the rhs is paulson@15409: directly 0, and @{term "Union C"} is also infinite, hence the lhs is also 0.*} 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" hoelzl@44937: proof cases hoelzl@44937: assume "finite C" hoelzl@44937: from setsum_UN_disjoint[OF this assms] hoelzl@44937: show ?thesis hoelzl@44937: by (simp add: SUP_def) hoelzl@44937: qed (force dest: finite_UnionD simp add: setsum_def) nipkow@15402: paulson@15409: (*But we can't get rid of finite A. If infinite, although the lhs is 0, paulson@15409: the rhs need not be, since SIGMA A B could still be finite.*) haftmann@35816: lemma (in comm_monoid_add) setsum_Sigma: haftmann@35816: assumes "finite A" and "ALL x:A. finite (B x)" haftmann@35816: shows "(\x\A. (\y\B x. f x y)) = (\(x,y)\(SIGMA x:A. B x). f x y)" haftmann@35816: proof - haftmann@35816: interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult) wenzelm@41550: from assms show ?thesis by (simp add: setsum_def fold_image_Sigma split_def) haftmann@35816: qed nipkow@15402: paulson@15409: text{*Here we can eliminate the finiteness assumptions, by cases.*} paulson@15409: lemma setsum_cartesian_product: paulson@17189: "(\x\A. (\y\B. f x y)) = (\(x,y) \ A <*> B. f x y)" paulson@15409: apply (cases "finite A") paulson@15409: apply (cases "finite B") paulson@15409: apply (simp add: setsum_Sigma) paulson@15409: apply (cases "A={}", simp) nipkow@15543: apply (simp) paulson@15409: apply (auto simp add: setsum_def paulson@15409: dest: finite_cartesian_productD1 finite_cartesian_productD2) paulson@15409: done nipkow@15402: nipkow@48861: lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)" nipkow@48861: by (fact setsum.F_fun_f) nipkow@15402: nipkow@48893: lemma setsum_Un_zero: nipkow@48893: "\ finite S; finite T; \x \ S\T. f x = 0 \ \ nipkow@48893: setsum f (S \ T) = setsum f S + setsum f T" nipkow@48893: by(fact setsum.F_Un_neutral) nipkow@48893: nipkow@48893: 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 nipkow@48893: by (auto simp add: H[symmetric] intro: setsum_Un_zero[OF fTF(1) fUF, of f]) nipkow@48893: qed nipkow@48893: nipkow@15402: nipkow@15402: subsubsection {* Properties in more restricted classes of structures *} nipkow@15402: nipkow@15402: lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a" nipkow@28853: apply (case_tac "finite A") nipkow@28853: prefer 2 apply (simp add: setsum_def) nipkow@28853: apply (erule rev_mp) nipkow@28853: apply (erule finite_induct, auto) nipkow@28853: done nipkow@15402: nipkow@15402: lemma setsum_eq_0_iff [simp]: nipkow@15402: "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))" nipkow@28853: by (induct set: finite) auto nipkow@15402: nipkow@30859: lemma setsum_eq_Suc0_iff: "finite A \ nipkow@30859: (setsum f A = Suc 0) = (EX a:A. f a = Suc 0 & (ALL b:A. a\b \ f b = 0))" nipkow@30859: apply(erule finite_induct) nipkow@30859: apply (auto simp add:add_is_1) nipkow@30859: done nipkow@30859: nipkow@30859: lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]] nipkow@30859: nipkow@15402: lemma setsum_Un_nat: "finite A ==> finite B ==> nipkow@28853: (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)" nipkow@15402: -- {* For the natural numbers, we have subtraction. *} nipkow@29667: by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps) 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@35816: lemma (in comm_monoid_add) setsum_eq_general_reverses: chaieb@30260: assumes fS: "finite S" and fT: "finite T" chaieb@30260: and kh: "\y. y \ T \ k y \ S \ h (k y) = y" chaieb@30260: and hk: "\x. x \ S \ h x \ T \ k (h x) = x \ g (h x) = f x" chaieb@30260: shows "setsum f S = setsum g T" haftmann@35816: proof - haftmann@35816: interpret comm_monoid_mult "op +" 0 by (fact comm_monoid_mult) haftmann@35816: show ?thesis chaieb@30260: apply (simp add: setsum_def fS fT) haftmann@35816: apply (rule fold_image_eq_general_inverses) haftmann@35816: apply (rule fS) chaieb@30260: apply (erule kh) chaieb@30260: apply (erule hk) chaieb@30260: done haftmann@35816: qed chaieb@30260: nipkow@15402: lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) = nipkow@28853: (if a:A then setsum f A - f a else setsum f A)" nipkow@28853: apply (case_tac "finite A") nipkow@28853: prefer 2 apply (simp add: setsum_def) nipkow@28853: apply (erule finite_induct) nipkow@28853: apply (auto simp add: insert_Diff_if) nipkow@28853: apply (drule_tac a = a in mk_disjoint_insert, auto) nipkow@28853: done nipkow@15402: 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@28853: lemma setsum_diff1'[rule_format]: nipkow@28853: "finite A \ a \ A \ (\ x \ A. f x) = f a + (\ x \ (A - {a}). f x)" nipkow@28853: apply (erule finite_induct[where F=A and P="% A. (a \ A \ (\ x \ A. f x) = f a + (\ x \ (A - {a}). f x))"]) nipkow@28853: apply (auto simp add: insert_Diff_if add_ac) nipkow@28853: done obua@15552: nipkow@31438: lemma setsum_diff1_ring: assumes "finite A" "a \ A" nipkow@31438: shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)" nipkow@31438: unfolding setsum_diff1'[OF assms] by auto nipkow@31438: nipkow@15402: (* By Jeremy Siek: *) nipkow@15402: nipkow@15402: lemma setsum_diff_nat: nipkow@28853: assumes "finite B" and "B \ A" nipkow@28853: shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)" nipkow@28853: using assms wenzelm@19535: proof induct nipkow@15402: show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp nipkow@15402: next nipkow@15402: fix F x assume finF: "finite F" and xnotinF: "x \ F" nipkow@15402: and xFinA: "insert x F \ A" nipkow@15402: and IH: "F \ A \ setsum f (A - F) = setsum f A - setsum f F" nipkow@15402: from xnotinF xFinA have xinAF: "x \ (A - F)" by simp nipkow@15402: from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x" nipkow@15402: by (simp add: setsum_diff1_nat) nipkow@15402: from xFinA have "F \ A" by simp nipkow@15402: with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp nipkow@15402: with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x" nipkow@15402: by simp nipkow@15402: from xnotinF have "A - insert x F = (A - F) - {x}" by auto nipkow@15402: with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x" nipkow@15402: by simp nipkow@15402: from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp nipkow@15402: with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" nipkow@15402: by simp nipkow@15402: thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp nipkow@15402: qed nipkow@15402: 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 nipkow@15402: case False nipkow@15402: thus ?thesis nipkow@15402: by (simp add: setsum_def) 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 nipkow@15535: case False thus ?thesis by (simp add: setsum_def) 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 nipkow@15535: case False thus ?thesis by (simp add: setsum_def) 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 nipkow@15535: case False thus ?thesis by (simp add: setsum_def) 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 nipkow@15535: case False thus ?thesis by (simp add: setsum_def) 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 nipkow@15402: case (insert x A) thus ?case by (simp add: right_distrib) nipkow@15402: qed nipkow@15402: next nipkow@15402: case False thus ?thesis by (simp add: setsum_def) 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 ballarin@17149: case (insert x A) thus ?case by (simp add: left_distrib) ballarin@17149: qed ballarin@17149: next ballarin@17149: case False thus ?thesis by (simp add: setsum_def) 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 ballarin@17149: case False thus ?thesis by (simp add: setsum_def) 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 nipkow@15535: case False thus ?thesis by (simp add: setsum_def) 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 nipkow@15535: case False thus ?thesis by (simp add: setsum_def) 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 nipkow@15539: case False thus ?thesis by (simp add: setsum_def) nipkow@15539: qed nipkow@15539: nipkow@31080: lemma setsum_Plus: nipkow@31080: fixes A :: "'a set" and B :: "'b set" nipkow@31080: assumes fin: "finite A" "finite B" nipkow@31080: shows "setsum f (A <+> B) = setsum (f \ Inl) A + setsum (f \ Inr) B" nipkow@31080: proof - nipkow@31080: have "A <+> B = Inl ` A \ Inr ` B" by auto nipkow@31080: moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)" nipkow@40786: by auto nipkow@31080: moreover have "Inl ` A \ Inr ` B = ({} :: ('a + 'b) set)" by auto nipkow@31080: moreover have "inj_on (Inl :: 'a \ 'a + 'b) A" "inj_on (Inr :: 'b \ 'a + 'b) B" by(auto intro: inj_onI) nipkow@31080: ultimately show ?thesis using fin by(simp add: setsum_Un_disjoint setsum_reindex) nipkow@31080: qed nipkow@31080: nipkow@31080: ballarin@17149: text {* Commuting outer and inner summation *} ballarin@17149: ballarin@17149: lemma setsum_commute: ballarin@17149: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" ballarin@17149: proof (simp add: setsum_cartesian_product) paulson@17189: have "(\(x,y) \ A <*> B. f x y) = paulson@17189: (\(y,x) \ (%(i, j). (j, i)) ` (A \ B). f x y)" ballarin@17149: (is "?s = _") ballarin@17149: apply (simp add: setsum_reindex [where f = "%(i, j). (j, i)"] swap_inj_on) ballarin@17149: apply (simp add: split_def) ballarin@17149: done paulson@17189: also have "... = (\(y,x)\B \ A. f x y)" ballarin@17149: (is "_ = ?t") ballarin@17149: apply (simp add: swap_product) ballarin@17149: done ballarin@17149: finally show "?s = ?t" . ballarin@17149: qed 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@35722: lemma setsum_constant [simp]: "(\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@35722: 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@35722: case False thus ?thesis by (simp add: setsum_def) haftmann@35722: qed haftmann@35722: haftmann@35722: haftmann@35722: subsubsection {* Cardinality as special case of @{const setsum} *} haftmann@35722: haftmann@35722: lemma card_eq_setsum: haftmann@35722: "card A = setsum (\x. 1) A" haftmann@35722: by (simp only: card_def setsum_def) 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: text{*The image of a finite set can be expressed using @{term fold_image}.*} haftmann@35722: lemma image_eq_fold_image: haftmann@35722: "finite A ==> f ` A = fold_image (op Un) (%x. {f x}) {} A" haftmann@35722: proof (induct rule: finite_induct) haftmann@35722: case empty then show ?case by simp haftmann@35722: next haftmann@35722: interpret ab_semigroup_mult "op Un" haftmann@35722: proof qed auto haftmann@35722: case insert haftmann@35722: then show ?case by simp haftmann@35722: qed 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@35816: definition (in comm_monoid_mult) setprod :: "('b \ 'a) => 'b set => 'a" where haftmann@35816: "setprod f A = (if finite A then fold_image (op *) f 1 A else 1)" haftmann@35816: huffman@35938: sublocale comm_monoid_mult < setprod!: comm_monoid_big "op *" 1 setprod proof haftmann@35816: qed (fact setprod_def) nipkow@15402: wenzelm@19535: abbreviation wenzelm@21404: Setprod ("\_" [1000] 999) where wenzelm@19535: "\A == setprod (%x. x) A" wenzelm@19535: 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@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: nipkow@28853: "inj_on f B ==> setprod h (f ` B) = setprod (h \ f) B" nipkow@48849: by(auto simp: setprod_def fold_image_reindex o_def dest!:finite_imageD) nipkow@15402: nipkow@15402: lemma setprod_reindex_id: "inj_on f B ==> setprod f B = setprod id (f ` B)" nipkow@15402: by (auto simp add: 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" nipkow@48849: by(fact setprod.F_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" nipkow@48849: by(fact setprod.strong_F_cong) berghofe@16632: nipkow@15402: lemma setprod_reindex_cong: "inj_on f A ==> nipkow@15402: B = f ` A ==> g = h \ f ==> setprod h B = setprod g A" nipkow@28853: by (frule setprod_reindex, simp) nipkow@15402: chaieb@29674: lemma strong_setprod_reindex_cong: assumes i: "inj_on f A" chaieb@29674: and B: "B = f ` A" and eq: "\x. x \ A \ g x = (h \ f) x" chaieb@29674: shows "setprod h B = setprod g A" chaieb@29674: proof- chaieb@29674: have "setprod h B = setprod (h o f) A" chaieb@29674: by (simp add: B setprod_reindex[OF i, of h]) chaieb@29674: then show ?thesis apply simp chaieb@29674: apply (rule setprod_cong) chaieb@29674: apply simp nipkow@30837: by (simp add: eq) chaieb@29674: qed chaieb@29674: nipkow@48893: lemma setprod_Un_one: "\ finite S; finite T; \x \ S\T. f x = 1 \ nipkow@48893: \ setprod f (S \ T) = setprod f S * setprod f T" nipkow@48893: by(fact setprod.F_Un_neutral) nipkow@15402: nipkow@48821: lemmas setprod_1 = setprod.F_neutral nipkow@48821: lemmas setprod_1' = setprod.F_neutral' nipkow@15402: 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" nipkow@48849: 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" nipkow@48849: 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" nipkow@48849: by(fact setprod.F_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" nipkow@48849: by(fact setprod.F_mono_neutral_left) nipkow@30837: nipkow@48849: lemmas setprod_mono_one_right = setprod.F_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" nipkow@48849: by(fact setprod.F_mono_neutral_cong_left) nipkow@48849: nipkow@48849: lemmas setprod_mono_one_cong_right = setprod.F_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)" nipkow@48849: by(fact setprod.F_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)" nipkow@48849: by(fact setprod.F_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" wenzelm@41550: by (simp add: setprod_def fold_image_UN_disjoint) nipkow@15402: nipkow@15402: lemma setprod_Union_disjoint: hoelzl@44937: assumes "\A\C. finite A" "\A\C. \B\C. A \ B \ A Int B = {}" hoelzl@44937: shows "setprod f (Union C) = setprod (setprod f) C" hoelzl@44937: proof cases hoelzl@44937: assume "finite C" hoelzl@44937: from setprod_UN_disjoint[OF this assms] hoelzl@44937: show ?thesis hoelzl@44937: by (simp add: SUP_def) hoelzl@44937: qed (force dest: finite_UnionD simp add: setprod_def) 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)" wenzelm@41550: by(simp add:setprod_def fold_image_Sigma split_def) nipkow@15402: paulson@15409: text{*Here we can eliminate the finiteness assumptions, by cases.*} paulson@15409: lemma setprod_cartesian_product: paulson@17189: "(\x\A. (\y\ B. f x y)) = (\(x,y)\(A <*> B). f x y)" paulson@15409: apply (cases "finite A") paulson@15409: apply (cases "finite B") paulson@15409: apply (simp add: setprod_Sigma) paulson@15409: apply (cases "A={}", simp) nipkow@48849: apply (simp) paulson@15409: apply (auto simp add: setprod_def paulson@15409: dest: finite_cartesian_productD1 finite_cartesian_productD2) paulson@15409: done nipkow@15402: nipkow@48861: lemma setprod_timesf: "setprod (%x. f x * g x) A = (setprod f A * setprod g A)" nipkow@48861: by (fact setprod.F_fun_f) nipkow@15402: nipkow@15402: nipkow@15402: subsubsection {* Properties in more restricted classes of structures *} nipkow@15402: nipkow@15402: lemma setprod_eq_1_iff [simp]: nipkow@28853: "finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))" nipkow@28853: by (induct set: finite) auto 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: 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@30843: lemma setprod_zero_iff[simp]: "finite A ==> nipkow@30843: (setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors})) = nipkow@30843: (EX x: A. f x = 0)" nipkow@30843: by (erule finite_induct, auto simp:no_zero_divisors) nipkow@30843: nipkow@30843: lemma setprod_pos_nat: nipkow@30843: "finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0" nipkow@30843: using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric]) nipkow@15402: nipkow@30863: lemma setprod_pos_nat_iff[simp]: nipkow@30863: "finite S ==> (setprod f S > 0) = (ALL x : S. f x > (0::nat))" nipkow@30863: using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric]) nipkow@30863: nipkow@15402: lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \ 0) ==> nipkow@28853: (setprod f (A Un B) :: 'a ::{field}) nipkow@28853: = setprod f A * setprod f B / setprod f (A Int B)" nipkow@30843: by (subst setprod_Un_Int [symmetric], auto) 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@31017: 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: chaieb@29674: haftmann@35816: subsection {* Versions of @{const inf} and @{const sup} on non-empty sets *} haftmann@35816: haftmann@35816: no_notation times (infixl "*" 70) haftmann@35816: no_notation Groups.one ("1") haftmann@35816: haftmann@35816: locale semilattice_big = semilattice + haftmann@35816: fixes F :: "'a set \ 'a" haftmann@35816: assumes F_eq: "finite A \ F A = fold1 (op *) A" haftmann@35816: haftmann@35816: sublocale semilattice_big < folding_one_idem proof haftmann@35816: qed (simp_all add: F_eq) haftmann@35816: haftmann@35816: notation times (infixl "*" 70) haftmann@35816: notation Groups.one ("1") haftmann@22917: haftmann@35816: context lattice haftmann@35816: begin haftmann@35816: haftmann@35816: definition Inf_fin :: "'a set \ 'a" ("\\<^bsub>fin\<^esub>_" [900] 900) where haftmann@35816: "Inf_fin = fold1 inf" haftmann@35816: haftmann@35816: definition Sup_fin :: "'a set \ 'a" ("\\<^bsub>fin\<^esub>_" [900] 900) where haftmann@35816: "Sup_fin = fold1 sup" haftmann@35816: haftmann@35816: end haftmann@35816: haftmann@35816: sublocale lattice < Inf_fin!: semilattice_big inf Inf_fin proof haftmann@35816: qed (simp add: Inf_fin_def) haftmann@35816: haftmann@35816: sublocale lattice < Sup_fin!: semilattice_big sup Sup_fin proof haftmann@35816: qed (simp add: Sup_fin_def) haftmann@22917: haftmann@35028: context semilattice_inf haftmann@26041: begin haftmann@26041: haftmann@36635: lemma ab_semigroup_idem_mult_inf: haftmann@36635: "class.ab_semigroup_idem_mult inf" haftmann@35816: proof qed (rule inf_assoc inf_commute inf_idem)+ haftmann@35816: haftmann@46033: lemma fold_inf_insert[simp]: "finite A \ Finite_Set.fold inf b (insert a A) = inf a (Finite_Set.fold inf b A)" haftmann@42871: by(rule comp_fun_idem.fold_insert_idem[OF ab_semigroup_idem_mult.comp_fun_idem[OF ab_semigroup_idem_mult_inf]]) haftmann@35816: haftmann@46033: lemma inf_le_fold_inf: "finite A \ ALL a:A. b \ a \ inf b c \ Finite_Set.fold inf c A" haftmann@35816: by (induct pred: finite) (auto intro: le_infI1) haftmann@35816: haftmann@46033: lemma fold_inf_le_inf: "finite A \ a \ A \ Finite_Set.fold inf b A \ inf a b" haftmann@35816: proof(induct arbitrary: a pred:finite) haftmann@35816: case empty thus ?case by simp haftmann@35816: next haftmann@35816: case (insert x A) haftmann@35816: show ?case haftmann@35816: proof cases haftmann@35816: assume "A = {}" thus ?thesis using insert by simp haftmann@35816: next haftmann@35816: assume "A \ {}" thus ?thesis using insert by (auto intro: le_infI2) haftmann@35816: qed haftmann@35816: qed haftmann@35816: haftmann@26041: lemma below_fold1_iff: haftmann@26041: assumes "finite A" "A \ {}" haftmann@26041: shows "x \ fold1 inf A \ (\a\A. x \ a)" haftmann@26041: proof - haftmann@29509: interpret ab_semigroup_idem_mult inf haftmann@26041: by (rule ab_semigroup_idem_mult_inf) haftmann@26041: show ?thesis using assms by (induct rule: finite_ne_induct) simp_all haftmann@26041: qed haftmann@26041: haftmann@26041: lemma fold1_belowI: haftmann@26757: assumes "finite A" haftmann@26041: and "a \ A" haftmann@26041: shows "fold1 inf A \ a" haftmann@26757: proof - haftmann@26757: from assms have "A \ {}" by auto haftmann@26757: from `finite A` `A \ {}` `a \ A` show ?thesis haftmann@26757: proof (induct rule: finite_ne_induct) haftmann@26757: case singleton thus ?case by simp haftmann@26041: next haftmann@29509: interpret ab_semigroup_idem_mult inf haftmann@26757: by (rule ab_semigroup_idem_mult_inf) haftmann@26757: case (insert x F) haftmann@26757: from insert(5) have "a = x \ a \ F" by simp haftmann@26757: thus ?case haftmann@26757: proof haftmann@26757: assume "a = x" thus ?thesis using insert nipkow@29667: by (simp add: mult_ac) haftmann@26757: next haftmann@26757: assume "a \ F" haftmann@26757: hence bel: "fold1 inf F \ a" by (rule insert) haftmann@26757: have "inf (fold1 inf (insert x F)) a = inf x (inf (fold1 inf F) a)" nipkow@29667: using insert by (simp add: mult_ac) haftmann@26757: also have "inf (fold1 inf F) a = fold1 inf F" haftmann@26757: using bel by (auto intro: antisym) haftmann@26757: also have "inf x \ = fold1 inf (insert x F)" nipkow@29667: using insert by (simp add: mult_ac) haftmann@26757: finally have aux: "inf (fold1 inf (insert x F)) a = fold1 inf (insert x F)" . haftmann@26757: moreover have "inf (fold1 inf (insert x F)) a \ a" by simp haftmann@26757: ultimately show ?thesis by simp haftmann@26757: qed haftmann@26041: qed haftmann@26041: qed haftmann@26041: haftmann@26041: end haftmann@26041: haftmann@35816: context semilattice_sup haftmann@22917: begin haftmann@22917: haftmann@36635: lemma ab_semigroup_idem_mult_sup: "class.ab_semigroup_idem_mult sup" haftmann@35816: by (rule semilattice_inf.ab_semigroup_idem_mult_inf)(rule dual_semilattice) haftmann@35816: haftmann@46033: lemma fold_sup_insert[simp]: "finite A \ Finite_Set.fold sup b (insert a A) = sup a (Finite_Set.fold sup b A)" haftmann@35816: by(rule semilattice_inf.fold_inf_insert)(rule dual_semilattice) haftmann@22917: haftmann@46033: lemma fold_sup_le_sup: "finite A \ ALL a:A. a \ b \ Finite_Set.fold sup c A \ sup b c" haftmann@35816: by(rule semilattice_inf.inf_le_fold_inf)(rule dual_semilattice) haftmann@35816: haftmann@46033: lemma sup_le_fold_sup: "finite A \ a \ A \ sup a b \ Finite_Set.fold sup b A" haftmann@35816: by(rule semilattice_inf.fold_inf_le_inf)(rule dual_semilattice) haftmann@35816: haftmann@35816: end haftmann@35816: haftmann@35816: context lattice haftmann@35816: begin haftmann@25062: wenzelm@31916: lemma Inf_le_Sup [simp]: "\ finite A; A \ {} \ \ \\<^bsub>fin\<^esub>A \ \\<^bsub>fin\<^esub>A" haftmann@24342: apply(unfold Sup_fin_def Inf_fin_def) 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@26757: apply(erule (1) fold1_belowI) haftmann@35028: apply(erule (1) semilattice_inf.fold1_belowI [OF dual_semilattice]) nipkow@15500: done nipkow@15500: haftmann@24342: lemma sup_Inf_absorb [simp]: wenzelm@31916: "finite A \ a \ A \ sup a (\\<^bsub>fin\<^esub>A) = a" nipkow@15512: apply(subst sup_commute) haftmann@26041: apply(simp add: Inf_fin_def sup_absorb2 fold1_belowI) nipkow@15504: done nipkow@15504: haftmann@24342: lemma inf_Sup_absorb [simp]: wenzelm@31916: "finite A \ a \ A \ inf a (\\<^bsub>fin\<^esub>A) = a" haftmann@26041: by (simp add: Sup_fin_def inf_absorb1 haftmann@35028: semilattice_inf.fold1_belowI [OF dual_semilattice]) 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@31916: shows "sup x (\\<^bsub>fin\<^esub>A) = \\<^bsub>fin\<^esub>{sup x a|a. a \ A}" haftmann@26041: proof - haftmann@29509: interpret ab_semigroup_idem_mult inf haftmann@26041: by (rule ab_semigroup_idem_mult_inf) haftmann@26041: from assms show ?thesis haftmann@26041: by (simp add: Inf_fin_def image_def haftmann@26041: hom_fold1_commute [where h="sup x", OF sup_inf_distrib1]) berghofe@26792: (rule arg_cong [where f="fold1 inf"], blast) haftmann@26041: qed nipkow@18423: haftmann@24342: lemma sup_Inf2_distrib: haftmann@24342: assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" wenzelm@31916: shows "sup (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B) = \\<^bsub>fin\<^esub>{sup a b|a b. a \ A \ b \ B}" haftmann@24342: using A proof (induct rule: finite_ne_induct) nipkow@15500: case singleton thus ?case wenzelm@41550: by (simp add: sup_Inf1_distrib [OF B]) nipkow@15500: next haftmann@29509: interpret ab_semigroup_idem_mult inf haftmann@26041: by (rule ab_semigroup_idem_mult_inf) nipkow@15500: case (insert x A) haftmann@25062: have finB: "finite {sup x b |b. b \ B}" haftmann@25062: 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@31916: have "sup (\\<^bsub>fin\<^esub>(insert x A)) (\\<^bsub>fin\<^esub>B) = sup (inf x (\\<^bsub>fin\<^esub>A)) (\\<^bsub>fin\<^esub>B)" wenzelm@41550: using insert by simp wenzelm@31916: also have "\ = inf (sup x (\\<^bsub>fin\<^esub>B)) (sup (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B))" by(rule sup_inf_distrib2) wenzelm@31916: also have "\ = inf (\\<^bsub>fin\<^esub>{sup x b|b. b \ B}) (\\<^bsub>fin\<^esub>{sup a b|a b. a \ A \ b \ B})" nipkow@15500: using insert by(simp add:sup_Inf1_distrib[OF B]) wenzelm@31916: also have "\ = \\<^bsub>fin\<^esub>({sup x b |b. b \ B} \ {sup a b |a b. a \ A \ b \ B})" wenzelm@31916: (is "_ = \\<^bsub>fin\<^esub>?M") nipkow@15500: using B insert haftmann@26041: by (simp add: Inf_fin_def fold1_Un2 [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@31916: shows "inf x (\\<^bsub>fin\<^esub>A) = \\<^bsub>fin\<^esub>{inf x a|a. a \ A}" haftmann@26041: proof - haftmann@29509: interpret ab_semigroup_idem_mult sup haftmann@26041: by (rule ab_semigroup_idem_mult_sup) haftmann@26041: from assms show ?thesis haftmann@26041: by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1]) berghofe@26792: (rule arg_cong [where f="fold1 sup"], blast) haftmann@26041: qed nipkow@18423: haftmann@24342: lemma inf_Sup2_distrib: haftmann@24342: assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" wenzelm@31916: shows "inf (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B) = \\<^bsub>fin\<^esub>{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 haftmann@29509: interpret ab_semigroup_idem_mult sup haftmann@26041: by (rule ab_semigroup_idem_mult_sup) wenzelm@31916: have "inf (\\<^bsub>fin\<^esub>(insert x A)) (\\<^bsub>fin\<^esub>B) = inf (sup x (\\<^bsub>fin\<^esub>A)) (\\<^bsub>fin\<^esub>B)" wenzelm@41550: using insert by simp wenzelm@31916: also have "\ = sup (inf x (\\<^bsub>fin\<^esub>B)) (inf (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B))" by(rule inf_sup_distrib2) wenzelm@31916: also have "\ = sup (\\<^bsub>fin\<^esub>{inf x b|b. b \ B}) (\\<^bsub>fin\<^esub>{inf a b|a b. a \ A \ b \ B})" nipkow@18423: using insert by(simp add:inf_Sup1_distrib[OF B]) wenzelm@31916: also have "\ = \\<^bsub>fin\<^esub>({inf x b |b. b \ B} \ {inf a b |a b. a \ A \ b \ B})" wenzelm@31916: (is "_ = \\<^bsub>fin\<^esub>?M") nipkow@18423: using B insert haftmann@26041: by (simp add: Sup_fin_def fold1_Un2 [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 \ {}" haftmann@35719: shows "\\<^bsub>fin\<^esub>A = Inf A" haftmann@35719: proof - haftmann@35719: interpret ab_semigroup_idem_mult inf haftmann@35719: by (rule ab_semigroup_idem_mult_inf) noschinl@44918: from `A \ {}` obtain b B where "A = {b} \ B" by auto haftmann@35719: moreover with `finite A` have "finite B" by simp noschinl@44918: ultimately show ?thesis noschinl@44918: by (simp add: Inf_fin_def fold1_eq_fold_idem inf_Inf_fold_inf [symmetric]) haftmann@35719: qed haftmann@35719: haftmann@35719: lemma Sup_fin_Sup: haftmann@35719: assumes "finite A" and "A \ {}" haftmann@35719: shows "\\<^bsub>fin\<^esub>A = Sup A" haftmann@35719: proof - haftmann@35719: interpret ab_semigroup_idem_mult sup haftmann@35719: by (rule ab_semigroup_idem_mult_sup) noschinl@44918: from `A \ {}` obtain b B where "A = {b} \ B" by auto haftmann@35719: moreover with `finite A` have "finite B" by simp haftmann@35719: ultimately show ?thesis haftmann@35719: by (simp add: Sup_fin_def fold1_eq_fold_idem sup_Sup_fold_sup [symmetric]) haftmann@35719: qed haftmann@35719: haftmann@35719: end haftmann@35719: haftmann@22917: haftmann@35816: subsection {* Versions of @{const min} and @{const max} on non-empty sets *} haftmann@35816: haftmann@35816: definition (in linorder) Min :: "'a set \ 'a" where haftmann@35816: "Min = fold1 min" haftmann@22917: haftmann@35816: definition (in linorder) Max :: "'a set \ 'a" where haftmann@35816: "Max = fold1 max" haftmann@35816: haftmann@35816: sublocale linorder < Min!: semilattice_big min Min proof haftmann@35816: qed (simp add: Min_def) haftmann@35816: haftmann@35816: sublocale linorder < Max!: semilattice_big max Max proof haftmann@35816: qed (simp add: Max_def) haftmann@22917: haftmann@24342: context linorder haftmann@22917: begin haftmann@22917: haftmann@35816: lemmas Min_singleton = Min.singleton haftmann@35816: lemmas Max_singleton = Max.singleton haftmann@35816: haftmann@35816: lemma Min_insert: haftmann@35816: assumes "finite A" and "A \ {}" haftmann@35816: shows "Min (insert x A) = min x (Min A)" haftmann@35816: using assms by simp haftmann@35816: haftmann@35816: lemma Max_insert: haftmann@35816: assumes "finite A" and "A \ {}" haftmann@35816: shows "Max (insert x A) = max x (Max A)" haftmann@35816: using assms by simp haftmann@35816: haftmann@35816: lemma Min_Un: haftmann@35816: assumes "finite A" and "A \ {}" and "finite B" and "B \ {}" haftmann@35816: shows "Min (A \ B) = min (Min A) (Min B)" haftmann@35816: using assms by (rule Min.union_idem) haftmann@35816: haftmann@35816: lemma Max_Un: haftmann@35816: assumes "finite A" and "A \ {}" and "finite B" and "B \ {}" haftmann@35816: shows "Max (A \ B) = max (Max A) (Max B)" haftmann@35816: using assms by (rule Max.union_idem) haftmann@35816: haftmann@35816: lemma hom_Min_commute: haftmann@35816: assumes "\x y. h (min x y) = min (h x) (h y)" haftmann@35816: and "finite N" and "N \ {}" haftmann@35816: shows "h (Min N) = Min (h ` N)" haftmann@35816: using assms by (rule Min.hom_commute) haftmann@35816: haftmann@35816: lemma hom_Max_commute: haftmann@35816: assumes "\x y. h (max x y) = max (h x) (h y)" haftmann@35816: and "finite N" and "N \ {}" haftmann@35816: shows "h (Max N) = Max (h ` N)" haftmann@35816: using assms by (rule Max.hom_commute) haftmann@35816: haftmann@26041: lemma ab_semigroup_idem_mult_min: haftmann@36635: "class.ab_semigroup_idem_mult min" haftmann@28823: proof qed (auto simp add: min_def) haftmann@26041: haftmann@26041: lemma ab_semigroup_idem_mult_max: haftmann@36635: "class.ab_semigroup_idem_mult max" haftmann@28823: proof qed (auto simp add: max_def) haftmann@26041: haftmann@26041: lemma max_lattice: krauss@44845: "class.semilattice_inf max (op \) (op >)" haftmann@32203: by (fact min_max.dual_semilattice) haftmann@26041: haftmann@26041: lemma dual_max: haftmann@26041: "ord.max (op \) = min" wenzelm@46904: by (auto simp add: ord.max_def min_def fun_eq_iff) haftmann@26041: haftmann@26041: lemma dual_min: haftmann@26041: "ord.min (op \) = max" wenzelm@46904: by (auto simp add: ord.min_def max_def fun_eq_iff) haftmann@26041: haftmann@26041: lemma strict_below_fold1_iff: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "x < fold1 min A \ (\a\A. x < a)" haftmann@26041: proof - haftmann@29509: interpret ab_semigroup_idem_mult min haftmann@26041: by (rule ab_semigroup_idem_mult_min) haftmann@26041: from assms show ?thesis haftmann@26041: by (induct rule: finite_ne_induct) haftmann@26041: (simp_all add: fold1_insert) haftmann@26041: qed haftmann@26041: haftmann@26041: lemma fold1_below_iff: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "fold1 min A \ x \ (\a\A. a \ x)" haftmann@26041: proof - haftmann@29509: interpret ab_semigroup_idem_mult min haftmann@26041: by (rule ab_semigroup_idem_mult_min) haftmann@26041: from assms show ?thesis haftmann@26041: by (induct rule: finite_ne_induct) haftmann@26041: (simp_all add: fold1_insert min_le_iff_disj) haftmann@26041: qed haftmann@26041: haftmann@26041: lemma fold1_strict_below_iff: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "fold1 min A < x \ (\a\A. a < x)" haftmann@26041: proof - haftmann@29509: interpret ab_semigroup_idem_mult min haftmann@26041: by (rule ab_semigroup_idem_mult_min) haftmann@26041: from assms show ?thesis haftmann@26041: by (induct rule: finite_ne_induct) haftmann@26041: (simp_all add: fold1_insert min_less_iff_disj) haftmann@26041: qed haftmann@26041: haftmann@26041: lemma fold1_antimono: haftmann@26041: assumes "A \ {}" and "A \ B" and "finite B" haftmann@26041: shows "fold1 min B \ fold1 min A" haftmann@26041: proof cases haftmann@26041: assume "A = B" thus ?thesis by simp haftmann@26041: next haftmann@29509: interpret ab_semigroup_idem_mult min haftmann@26041: by (rule ab_semigroup_idem_mult_min) wenzelm@41550: assume neq: "A \ B" haftmann@26041: have B: "B = A \ (B-A)" using `A \ B` by blast haftmann@26041: have "fold1 min B = fold1 min (A \ (B-A))" by(subst B)(rule refl) haftmann@26041: also have "\ = min (fold1 min A) (fold1 min (B-A))" haftmann@26041: proof - haftmann@26041: have "finite A" by(rule finite_subset[OF `A \ B` `finite B`]) wenzelm@41550: moreover have "finite(B-A)" by(rule finite_Diff[OF `finite B`]) wenzelm@41550: moreover have "(B-A) \ {}" using assms neq by blast wenzelm@41550: moreover have "A Int (B-A) = {}" using assms by blast haftmann@26041: ultimately show ?thesis using `A \ {}` by (rule_tac fold1_Un) haftmann@26041: qed haftmann@26041: also have "\ \ fold1 min A" by (simp add: min_le_iff_disj) haftmann@26041: finally show ?thesis . haftmann@26041: qed haftmann@26041: paulson@24427: lemma Min_in [simp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "Min A \ A" haftmann@26041: proof - haftmann@29509: interpret ab_semigroup_idem_mult min haftmann@26041: by (rule ab_semigroup_idem_mult_min) nipkow@44890: from assms fold1_in show ?thesis by (fastforce simp: Min_def min_def) haftmann@26041: qed nipkow@15392: paulson@24427: lemma Max_in [simp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "Max A \ A" haftmann@26041: proof - haftmann@29509: interpret ab_semigroup_idem_mult max haftmann@26041: by (rule ab_semigroup_idem_mult_max) nipkow@44890: from assms fold1_in [of A] show ?thesis by (fastforce simp: Max_def max_def) haftmann@26041: qed haftmann@26041: haftmann@26041: lemma Min_le [simp]: haftmann@26757: assumes "finite A" and "x \ A" haftmann@26041: shows "Min A \ x" haftmann@32203: using assms by (simp add: Min_def min_max.fold1_belowI) haftmann@26041: haftmann@26041: lemma Max_ge [simp]: haftmann@26757: assumes "finite A" and "x \ A" haftmann@26041: shows "x \ Max A" huffman@44921: by (simp add: Max_def semilattice_inf.fold1_belowI [OF max_lattice] assms) haftmann@26041: blanchet@35828: lemma Min_ge_iff [simp, no_atp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "x \ Min A \ (\a\A. x \ a)" haftmann@32203: using assms by (simp add: Min_def min_max.below_fold1_iff) haftmann@26041: blanchet@35828: lemma Max_le_iff [simp, no_atp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "Max A \ x \ (\a\A. a \ x)" huffman@44921: by (simp add: Max_def semilattice_inf.below_fold1_iff [OF max_lattice] assms) haftmann@26041: blanchet@35828: lemma Min_gr_iff [simp, no_atp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "x < Min A \ (\a\A. x < a)" haftmann@32203: using assms by (simp add: Min_def strict_below_fold1_iff) haftmann@26041: blanchet@35828: lemma Max_less_iff [simp, no_atp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "Max A < x \ (\a\A. a < x)" huffman@44921: by (simp add: Max_def linorder.dual_max [OF dual_linorder] huffman@44921: linorder.strict_below_fold1_iff [OF dual_linorder] assms) nipkow@18493: blanchet@35828: lemma Min_le_iff [no_atp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "Min A \ x \ (\a\A. a \ x)" haftmann@32203: using assms by (simp add: Min_def fold1_below_iff) nipkow@15497: blanchet@35828: lemma Max_ge_iff [no_atp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "x \ Max A \ (\a\A. x \ a)" huffman@44921: by (simp add: Max_def linorder.dual_max [OF dual_linorder] huffman@44921: linorder.fold1_below_iff [OF dual_linorder] assms) haftmann@22917: blanchet@35828: lemma Min_less_iff [no_atp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "Min A < x \ (\a\A. a < x)" haftmann@32203: using assms by (simp add: Min_def fold1_strict_below_iff) haftmann@22917: blanchet@35828: lemma Max_gr_iff [no_atp]: haftmann@26041: assumes "finite A" and "A \ {}" haftmann@26041: shows "x < Max A \ (\a\A. x < a)" huffman@44921: by (simp add: Max_def linorder.dual_max [OF dual_linorder] huffman@44921: linorder.fold1_strict_below_iff [OF dual_linorder] assms) 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@26041: lemma Min_antimono: haftmann@26041: assumes "M \ N" and "M \ {}" and "finite N" haftmann@26041: shows "Min N \ Min M" haftmann@32203: using assms by (simp add: Min_def fold1_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" huffman@44921: by (simp add: Max_def linorder.dual_max [OF dual_linorder] huffman@44921: linorder.fold1_antimono [OF dual_linorder] assms) haftmann@22917: nipkow@32006: lemma finite_linorder_max_induct[consumes 1, case_names empty insert]: urbanc@36079: assumes fin: "finite A" urbanc@36079: and empty: "P {}" urbanc@36079: and insert: "(!!b A. finite A \ ALL a:A. a < b \ P A \ P(insert b A))" urbanc@36079: 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 urbanc@36079: 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 nipkow@44890: ultimately show "P A" using A insert_Diff_single step[OF `finite ?B`] by fastforce krauss@26748: qed krauss@26748: qed krauss@26748: nipkow@32006: lemma finite_linorder_min_induct[consumes 1, case_names empty insert]: nipkow@33434: "\finite A; P {}; \b A. \finite A; \a\A. b < a; P A\ \ P (insert b A)\ \ P A" nipkow@32006: by(rule linorder.finite_linorder_max_induct[OF dual_linorder]) nipkow@32006: 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@35034: "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@35034: "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@25571: end