wenzelm@12396: (* Title: HOL/Finite_Set.thy wenzelm@12396: ID: $Id$ wenzelm@12396: Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel paulson@14430: Additions by Jeremy Avigad in Feb 2004 nipkow@15376: nipkow@15402: Get rid of a couple of superfluous finiteness assumptions in lemmas nipkow@15402: about setsum and card - see FIXME. nipkow@15402: NB: the analogous lemmas for setprod should also be simplified! wenzelm@12396: *) wenzelm@12396: wenzelm@12396: header {* Finite sets *} wenzelm@12396: nipkow@15131: theory Finite_Set nipkow@15140: imports Divides Power Inductive nipkow@15131: begin wenzelm@12396: nipkow@15392: subsection {* Definition and basic properties *} wenzelm@12396: wenzelm@12396: consts Finites :: "'a set set" nipkow@13737: syntax nipkow@13737: finite :: "'a set => bool" nipkow@13737: translations nipkow@13737: "finite A" == "A : Finites" wenzelm@12396: wenzelm@12396: inductive Finites wenzelm@12396: intros wenzelm@12396: emptyI [simp, intro!]: "{} : Finites" wenzelm@12396: insertI [simp, intro!]: "A : Finites ==> insert a A : Finites" wenzelm@12396: wenzelm@12396: axclass finite \ type wenzelm@12396: finite: "finite UNIV" wenzelm@12396: nipkow@13737: lemma ex_new_if_finite: -- "does not depend on def of finite at all" wenzelm@14661: assumes "\ finite (UNIV :: 'a set)" and "finite A" wenzelm@14661: shows "\a::'a. a \ A" wenzelm@14661: proof - wenzelm@14661: from prems have "A \ UNIV" by blast wenzelm@14661: thus ?thesis by blast wenzelm@14661: qed wenzelm@12396: wenzelm@12396: lemma finite_induct [case_names empty insert, induct set: Finites]: wenzelm@12396: "finite F ==> nipkow@15327: P {} ==> (!!x F. finite F ==> x \ F ==> P F ==> P (insert x F)) ==> P F" wenzelm@12396: -- {* Discharging @{text "x \ F"} entails extra work. *} wenzelm@12396: proof - wenzelm@13421: assume "P {}" and nipkow@15327: insert: "!!x F. finite F ==> x \ F ==> P F ==> P (insert x F)" wenzelm@12396: assume "finite F" wenzelm@12396: thus "P F" wenzelm@12396: proof induct wenzelm@12396: show "P {}" . nipkow@15327: fix x F assume F: "finite F" and P: "P F" wenzelm@12396: show "P (insert x F)" wenzelm@12396: proof cases wenzelm@12396: assume "x \ F" wenzelm@12396: hence "insert x F = F" by (rule insert_absorb) wenzelm@12396: with P show ?thesis by (simp only:) wenzelm@12396: next wenzelm@12396: assume "x \ F" wenzelm@12396: from F this P show ?thesis by (rule insert) wenzelm@12396: qed wenzelm@12396: qed wenzelm@12396: qed wenzelm@12396: wenzelm@12396: lemma finite_subset_induct [consumes 2, case_names empty insert]: wenzelm@12396: "finite F ==> F \ A ==> nipkow@15327: P {} ==> (!!a F. finite F ==> a \ A ==> a \ F ==> P F ==> P (insert a F)) ==> wenzelm@12396: P F" wenzelm@12396: proof - wenzelm@13421: assume "P {}" and insert: nipkow@15327: "!!a F. finite F ==> a \ A ==> a \ F ==> P F ==> P (insert a F)" wenzelm@12396: assume "finite F" wenzelm@12396: thus "F \ A ==> P F" wenzelm@12396: proof induct wenzelm@12396: show "P {}" . nipkow@15327: fix x F assume "finite F" and "x \ F" wenzelm@12396: and P: "F \ A ==> P F" and i: "insert x F \ A" wenzelm@12396: show "P (insert x F)" wenzelm@12396: proof (rule insert) wenzelm@12396: from i show "x \ A" by blast wenzelm@12396: from i have "F \ A" by blast wenzelm@12396: with P show "P F" . wenzelm@12396: qed wenzelm@12396: qed wenzelm@12396: qed wenzelm@12396: nipkow@15392: text{* Finite sets are the images of initial segments of natural numbers: *} nipkow@15392: nipkow@15392: lemma finite_imp_nat_seg_image: nipkow@15392: assumes fin: "finite A" shows "\ (n::nat) f. A = f ` {i::nat. if. {} = f ` {i::nat. i < 0}" by(simp add:image_def) qed nipkow@15392: next nipkow@15392: case (insert a A) nipkow@15392: from insert.hyps obtain n f where "A = f ` {i::nat. i < n}" by blast nipkow@15392: hence "insert a A = (%i. if i finite A" nipkow@15392: proof (induct n) nipkow@15392: case 0 thus ?case by simp nipkow@15392: next nipkow@15392: case (Suc n) nipkow@15392: let ?B = "f ` {i. i < n}" nipkow@15392: have finB: "finite ?B" by(rule Suc.hyps[OF refl]) nipkow@15392: show ?case nipkow@15392: proof cases nipkow@15392: assume "\k(\ k (n::nat) f. A = f ` {i::nat. i finite G ==> finite (F Un G)" wenzelm@12396: -- {* The union of two finite sets is finite. *} wenzelm@12396: by (induct set: Finites) simp_all wenzelm@12396: wenzelm@12396: lemma finite_subset: "A \ B ==> finite B ==> finite A" wenzelm@12396: -- {* Every subset of a finite set is finite. *} wenzelm@12396: proof - wenzelm@12396: assume "finite B" wenzelm@12396: thus "!!A. A \ B ==> finite A" wenzelm@12396: proof induct wenzelm@12396: case empty wenzelm@12396: thus ?case by simp wenzelm@12396: next nipkow@15327: case (insert x F A) wenzelm@12396: have A: "A \ insert x F" and r: "A - {x} \ F ==> finite (A - {x})" . wenzelm@12396: show "finite A" wenzelm@12396: proof cases wenzelm@12396: assume x: "x \ A" wenzelm@12396: with A have "A - {x} \ F" by (simp add: subset_insert_iff) wenzelm@12396: with r have "finite (A - {x})" . wenzelm@12396: hence "finite (insert x (A - {x}))" .. wenzelm@12396: also have "insert x (A - {x}) = A" by (rule insert_Diff) wenzelm@12396: finally show ?thesis . wenzelm@12396: next wenzelm@12396: show "A \ F ==> ?thesis" . wenzelm@12396: assume "x \ A" wenzelm@12396: with A show "A \ F" by (simp add: subset_insert_iff) wenzelm@12396: qed wenzelm@12396: qed wenzelm@12396: qed wenzelm@12396: wenzelm@12396: lemma finite_Un [iff]: "finite (F Un G) = (finite F & finite G)" wenzelm@12396: by (blast intro: finite_subset [of _ "X Un Y", standard] finite_UnI) wenzelm@12396: wenzelm@12396: lemma finite_Int [simp, intro]: "finite F | finite G ==> finite (F Int G)" wenzelm@12396: -- {* The converse obviously fails. *} wenzelm@12396: by (blast intro: finite_subset) wenzelm@12396: wenzelm@12396: lemma finite_insert [simp]: "finite (insert a A) = finite A" wenzelm@12396: apply (subst insert_is_Un) paulson@14208: apply (simp only: finite_Un, blast) wenzelm@12396: done wenzelm@12396: nipkow@15281: lemma finite_Union[simp, intro]: nipkow@15281: "\ finite A; !!M. M \ A \ finite M \ \ finite(\A)" nipkow@15281: by (induct rule:finite_induct) simp_all nipkow@15281: wenzelm@12396: lemma finite_empty_induct: wenzelm@12396: "finite A ==> wenzelm@12396: P A ==> (!!a A. finite A ==> a:A ==> P A ==> P (A - {a})) ==> P {}" wenzelm@12396: proof - wenzelm@12396: assume "finite A" wenzelm@12396: and "P A" and "!!a A. finite A ==> a:A ==> P A ==> P (A - {a})" wenzelm@12396: have "P (A - A)" wenzelm@12396: proof - wenzelm@12396: fix c b :: "'a set" wenzelm@12396: presume c: "finite c" and b: "finite b" wenzelm@12396: and P1: "P b" and P2: "!!x y. finite y ==> x \ y ==> P y ==> P (y - {x})" wenzelm@12396: from c show "c \ b ==> P (b - c)" wenzelm@12396: proof induct wenzelm@12396: case empty wenzelm@12396: from P1 show ?case by simp wenzelm@12396: next nipkow@15327: case (insert x F) wenzelm@12396: have "P (b - F - {x})" wenzelm@12396: proof (rule P2) wenzelm@12396: from _ b show "finite (b - F)" by (rule finite_subset) blast wenzelm@12396: from insert show "x \ b - F" by simp wenzelm@12396: from insert show "P (b - F)" by simp wenzelm@12396: qed wenzelm@12396: also have "b - F - {x} = b - insert x F" by (rule Diff_insert [symmetric]) wenzelm@12396: finally show ?case . wenzelm@12396: qed wenzelm@12396: next wenzelm@12396: show "A \ A" .. wenzelm@12396: qed wenzelm@12396: thus "P {}" by simp wenzelm@12396: qed wenzelm@12396: wenzelm@12396: lemma finite_Diff [simp]: "finite B ==> finite (B - Ba)" wenzelm@12396: by (rule Diff_subset [THEN finite_subset]) wenzelm@12396: wenzelm@12396: lemma finite_Diff_insert [iff]: "finite (A - insert a B) = finite (A - B)" wenzelm@12396: apply (subst Diff_insert) wenzelm@12396: apply (case_tac "a : A - B") wenzelm@12396: apply (rule finite_insert [symmetric, THEN trans]) paulson@14208: apply (subst insert_Diff, simp_all) wenzelm@12396: done wenzelm@12396: wenzelm@12396: nipkow@15392: text {* Image and Inverse Image over Finite Sets *} paulson@13825: paulson@13825: lemma finite_imageI[simp]: "finite F ==> finite (h ` F)" paulson@13825: -- {* The image of a finite set is finite. *} paulson@13825: by (induct set: Finites) simp_all paulson@13825: paulson@14430: lemma finite_surj: "finite A ==> B <= f ` A ==> finite B" paulson@14430: apply (frule finite_imageI) paulson@14430: apply (erule finite_subset, assumption) paulson@14430: done paulson@14430: paulson@13825: lemma finite_range_imageI: paulson@13825: "finite (range g) ==> finite (range (%x. f (g x)))" paulson@14208: apply (drule finite_imageI, simp) paulson@13825: done paulson@13825: wenzelm@12396: lemma finite_imageD: "finite (f`A) ==> inj_on f A ==> finite A" wenzelm@12396: proof - wenzelm@12396: have aux: "!!A. finite (A - {}) = finite A" by simp wenzelm@12396: fix B :: "'a set" wenzelm@12396: assume "finite B" wenzelm@12396: thus "!!A. f`A = B ==> inj_on f A ==> finite A" wenzelm@12396: apply induct wenzelm@12396: apply simp wenzelm@12396: apply (subgoal_tac "EX y:A. f y = x & F = f ` (A - {y})") wenzelm@12396: apply clarify wenzelm@12396: apply (simp (no_asm_use) add: inj_on_def) paulson@14208: apply (blast dest!: aux [THEN iffD1], atomize) wenzelm@12396: apply (erule_tac V = "ALL A. ?PP (A)" in thin_rl) paulson@14208: apply (frule subsetD [OF equalityD2 insertI1], clarify) wenzelm@12396: apply (rule_tac x = xa in bexI) wenzelm@12396: apply (simp_all add: inj_on_image_set_diff) wenzelm@12396: done wenzelm@12396: qed (rule refl) wenzelm@12396: wenzelm@12396: paulson@13825: lemma inj_vimage_singleton: "inj f ==> f-`{a} \ {THE x. f x = a}" paulson@13825: -- {* The inverse image of a singleton under an injective function paulson@13825: is included in a singleton. *} paulson@14430: apply (auto simp add: inj_on_def) paulson@14430: apply (blast intro: the_equality [symmetric]) paulson@13825: done paulson@13825: paulson@13825: lemma finite_vimageI: "[|finite F; inj h|] ==> finite (h -` F)" paulson@13825: -- {* The inverse image of a finite set under an injective function paulson@13825: is finite. *} paulson@14430: apply (induct set: Finites, simp_all) paulson@14430: apply (subst vimage_insert) paulson@14430: apply (simp add: finite_Un finite_subset [OF inj_vimage_singleton]) paulson@13825: done paulson@13825: paulson@13825: nipkow@15392: text {* The finite UNION of finite sets *} wenzelm@12396: wenzelm@12396: lemma finite_UN_I: "finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (UN a:A. B a)" wenzelm@12396: by (induct set: Finites) simp_all wenzelm@12396: wenzelm@12396: text {* wenzelm@12396: Strengthen RHS to paulson@14430: @{prop "((ALL x:A. finite (B x)) & finite {x. x:A & B x \ {}})"}? wenzelm@12396: wenzelm@12396: We'd need to prove paulson@14430: @{prop "finite C ==> ALL A B. (UNION A B) <= C --> finite {x. x:A & B x \ {}}"} wenzelm@12396: by induction. *} wenzelm@12396: wenzelm@12396: lemma finite_UN [simp]: "finite A ==> finite (UNION A B) = (ALL x:A. finite (B x))" wenzelm@12396: by (blast intro: finite_UN_I finite_subset) wenzelm@12396: wenzelm@12396: nipkow@15392: text {* Sigma of finite sets *} wenzelm@12396: wenzelm@12396: lemma finite_SigmaI [simp]: wenzelm@12396: "finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (SIGMA a:A. B a)" wenzelm@12396: by (unfold Sigma_def) (blast intro!: finite_UN_I) wenzelm@12396: nipkow@15402: lemma finite_cartesian_product: "[| finite A; finite B |] ==> nipkow@15402: finite (A <*> B)" nipkow@15402: by (rule finite_SigmaI) nipkow@15402: wenzelm@12396: lemma finite_Prod_UNIV: wenzelm@12396: "finite (UNIV::'a set) ==> finite (UNIV::'b set) ==> finite (UNIV::('a * 'b) set)" wenzelm@12396: apply (subgoal_tac "(UNIV:: ('a * 'b) set) = Sigma UNIV (%x. UNIV)") wenzelm@12396: apply (erule ssubst) paulson@14208: apply (erule finite_SigmaI, auto) wenzelm@12396: done wenzelm@12396: wenzelm@12396: instance unit :: finite wenzelm@12396: proof wenzelm@12396: have "finite {()}" by simp wenzelm@12396: also have "{()} = UNIV" by auto wenzelm@12396: finally show "finite (UNIV :: unit set)" . wenzelm@12396: qed wenzelm@12396: wenzelm@12396: instance * :: (finite, finite) finite wenzelm@12396: proof wenzelm@12396: show "finite (UNIV :: ('a \ 'b) set)" wenzelm@12396: proof (rule finite_Prod_UNIV) wenzelm@12396: show "finite (UNIV :: 'a set)" by (rule finite) wenzelm@12396: show "finite (UNIV :: 'b set)" by (rule finite) wenzelm@12396: qed wenzelm@12396: qed wenzelm@12396: wenzelm@12396: nipkow@15392: text {* The powerset of a finite set *} wenzelm@12396: wenzelm@12396: lemma finite_Pow_iff [iff]: "finite (Pow A) = finite A" wenzelm@12396: proof wenzelm@12396: assume "finite (Pow A)" wenzelm@12396: with _ have "finite ((%x. {x}) ` A)" by (rule finite_subset) blast wenzelm@12396: thus "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp wenzelm@12396: next wenzelm@12396: assume "finite A" wenzelm@12396: thus "finite (Pow A)" wenzelm@12396: by induct (simp_all add: finite_UnI finite_imageI Pow_insert) wenzelm@12396: qed wenzelm@12396: nipkow@15392: nipkow@15392: lemma finite_UnionD: "finite(\A) \ finite A" nipkow@15392: by(blast intro: finite_subset[OF subset_Pow_Union]) nipkow@15392: nipkow@15392: wenzelm@12396: lemma finite_converse [iff]: "finite (r^-1) = finite r" wenzelm@12396: apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r") wenzelm@12396: apply simp wenzelm@12396: apply (rule iffI) wenzelm@12396: apply (erule finite_imageD [unfolded inj_on_def]) wenzelm@12396: apply (simp split add: split_split) wenzelm@12396: apply (erule finite_imageI) paulson@14208: apply (simp add: converse_def image_def, auto) wenzelm@12396: apply (rule bexI) wenzelm@12396: prefer 2 apply assumption wenzelm@12396: apply simp wenzelm@12396: done wenzelm@12396: paulson@14430: nipkow@15392: text {* \paragraph{Finiteness of transitive closure} (Thanks to Sidi nipkow@15392: Ehmety) *} wenzelm@12396: wenzelm@12396: lemma finite_Field: "finite r ==> finite (Field r)" wenzelm@12396: -- {* A finite relation has a finite field (@{text "= domain \ range"}. *} wenzelm@12396: apply (induct set: Finites) wenzelm@12396: apply (auto simp add: Field_def Domain_insert Range_insert) wenzelm@12396: done wenzelm@12396: wenzelm@12396: lemma trancl_subset_Field2: "r^+ <= Field r \ Field r" wenzelm@12396: apply clarify wenzelm@12396: apply (erule trancl_induct) wenzelm@12396: apply (auto simp add: Field_def) wenzelm@12396: done wenzelm@12396: wenzelm@12396: lemma finite_trancl: "finite (r^+) = finite r" wenzelm@12396: apply auto wenzelm@12396: prefer 2 wenzelm@12396: apply (rule trancl_subset_Field2 [THEN finite_subset]) wenzelm@12396: apply (rule finite_SigmaI) wenzelm@12396: prefer 3 berghofe@13704: apply (blast intro: r_into_trancl' finite_subset) wenzelm@12396: apply (auto simp add: finite_Field) wenzelm@12396: done wenzelm@12396: wenzelm@12396: nipkow@15392: subsection {* A fold functional for finite sets *} nipkow@15392: nipkow@15392: text {* The intended behaviour is nipkow@15392: @{text "fold f g e {x\<^isub>1, ..., x\<^isub>n} = f (g x\<^isub>1) (\ (f (g x\<^isub>n) e)\)"} nipkow@15392: if @{text f} is associative-commutative. For an application of @{text fold} nipkow@15392: se the definitions of sums and products over finite sets. nipkow@15392: *} nipkow@15392: nipkow@15392: consts nipkow@15392: foldSet :: "('a => 'a => 'a) => ('b => 'a) => 'a => ('b set \ 'a) set" nipkow@15392: nipkow@15392: inductive "foldSet f g e" nipkow@15392: intros nipkow@15392: emptyI [intro]: "({}, e) : foldSet f g e" nipkow@15392: insertI [intro]: "\ x \ A; (A, y) : foldSet f g e \ nipkow@15392: \ (insert x A, f (g x) y) : foldSet f g e" nipkow@15392: nipkow@15392: inductive_cases empty_foldSetE [elim!]: "({}, x) : foldSet f g e" nipkow@15392: nipkow@15392: constdefs nipkow@15392: fold :: "('a => 'a => 'a) => ('b => 'a) => 'a => 'b set => 'a" nipkow@15392: "fold f g e A == THE x. (A, x) : foldSet f g e" nipkow@15392: nipkow@15392: lemma Diff1_foldSet: nipkow@15392: "(A - {x}, y) : foldSet f g e ==> x: A ==> (A, f (g x) y) : foldSet f g e" nipkow@15392: by (erule insert_Diff [THEN subst], rule foldSet.intros, auto) nipkow@15392: nipkow@15392: lemma foldSet_imp_finite: "(A, x) : foldSet f g e ==> finite A" nipkow@15392: by (induct set: foldSet) auto nipkow@15392: nipkow@15392: lemma finite_imp_foldSet: "finite A ==> EX x. (A, x) : foldSet f g e" nipkow@15392: by (induct set: Finites) auto nipkow@15392: nipkow@15392: nipkow@15392: subsubsection {* Commutative monoids *} nipkow@15392: nipkow@15392: locale ACf = nipkow@15392: fixes f :: "'a => 'a => 'a" (infixl "\" 70) nipkow@15392: assumes commute: "x \ y = y \ x" nipkow@15392: and assoc: "(x \ y) \ z = x \ (y \ z)" nipkow@15392: nipkow@15392: locale ACe = ACf + nipkow@15392: fixes e :: 'a nipkow@15392: assumes ident [simp]: "x \ e = x" nipkow@15392: nipkow@15392: lemma (in ACf) left_commute: "x \ (y \ z) = y \ (x \ z)" nipkow@15392: proof - nipkow@15392: have "x \ (y \ z) = (y \ z) \ x" by (simp only: commute) nipkow@15392: also have "... = y \ (z \ x)" by (simp only: assoc) nipkow@15392: also have "z \ x = x \ z" by (simp only: commute) nipkow@15392: finally show ?thesis . nipkow@15392: qed nipkow@15392: nipkow@15392: lemmas (in ACf) AC = assoc commute left_commute nipkow@15392: nipkow@15392: lemma (in ACe) left_ident [simp]: "e \ x = x" nipkow@15392: proof - nipkow@15392: have "x \ e = x" by (rule ident) nipkow@15392: thus ?thesis by (subst commute) nipkow@15392: qed nipkow@15392: nipkow@15402: text{* Instantiation of locales: *} nipkow@15402: nipkow@15402: lemma ACf_add: "ACf (op + :: 'a::comm_monoid_add \ 'a \ 'a)" nipkow@15402: by(fastsimp intro: ACf.intro add_assoc add_commute) nipkow@15402: nipkow@15402: lemma ACe_add: "ACe (op +) (0::'a::comm_monoid_add)" nipkow@15402: by(fastsimp intro: ACe.intro ACe_axioms.intro ACf_add) nipkow@15402: nipkow@15402: nipkow@15402: lemma ACf_mult: "ACf (op * :: 'a::comm_monoid_mult \ 'a \ 'a)" nipkow@15402: by(fast intro: ACf.intro mult_assoc ab_semigroup_mult.mult_commute) nipkow@15402: nipkow@15402: lemma ACe_mult: "ACe (op *) (1::'a::comm_monoid_mult)" nipkow@15402: by(fastsimp intro: ACe.intro ACe_axioms.intro ACf_mult) nipkow@15402: nipkow@15402: nipkow@15392: subsubsection{*From @{term foldSet} to @{term fold}*} nipkow@15392: nipkow@15392: lemma (in ACf) foldSet_determ_aux: nipkow@15392: "!!A x x' h. \ A = h`{i::nat. i nipkow@15392: \ x' = x" nipkow@15392: proof (induct n) nipkow@15392: case 0 thus ?case by auto nipkow@15392: next nipkow@15392: case (Suc n) nipkow@15392: have IH: "!!A x x' h. \A = h`{i::nat. i foldSet f g e; (A,x') \ foldSet f g e\ nipkow@15392: \ x' = x" and card: "A = h`{i. i foldSet f g e" and Afoldy: "(A,x') \ foldSet f g e" . nipkow@15392: show ?case nipkow@15392: proof cases nipkow@15392: assume "EX k(EX k y)" nipkow@15392: and y: "(B,y) \ foldSet f g e" and notinB: "b \ B" nipkow@15392: hence A1: "A = insert b B" and x: "x = g b \ y" by auto nipkow@15392: show ?thesis nipkow@15392: proof (rule foldSet.cases[OF Afoldy]) nipkow@15392: assume "(A,x') = ({}, e)" nipkow@15392: thus ?thesis using A1 by auto nipkow@15392: next nipkow@15392: fix C c z nipkow@15392: assume eq2: "(A,x') = (insert c C, g c \ z)" nipkow@15392: and z: "(C,z) \ foldSet f g e" and notinC: "c \ C" nipkow@15392: hence A2: "A = insert c C" and x': "x' = g c \ z" by auto nipkow@15392: let ?h = "%i. if h i = b then h n else h i" nipkow@15392: have less: "B = ?h`{i. i ?r" nipkow@15392: proof nipkow@15392: fix u assume "u \ B" nipkow@15392: hence uinA: "u \ A" and unotb: "u \ b" using A1 notinB by blast+ nipkow@15392: then obtain i\<^isub>u where below: "i\<^isub>u < Suc n" and [simp]: "u = h i\<^isub>u" nipkow@15392: using card by(auto simp:image_def) nipkow@15392: show "u \ ?r" nipkow@15392: proof cases nipkow@15392: assume "i\<^isub>u < n" nipkow@15392: thus ?thesis using unotb by(fastsimp) nipkow@15392: next nipkow@15392: assume "\ i\<^isub>u < n" nipkow@15392: with below have [simp]: "i\<^isub>u = n" by arith nipkow@15392: obtain i\<^isub>k where i\<^isub>k: "i\<^isub>k < Suc n" and [simp]: "b = h i\<^isub>k" nipkow@15392: using A1 card by blast nipkow@15392: have "i\<^isub>k < n" nipkow@15392: proof (rule ccontr) nipkow@15392: assume "\ i\<^isub>k < n" nipkow@15392: hence "i\<^isub>k = n" using i\<^isub>k by arith nipkow@15392: thus False using unotb by simp nipkow@15392: qed nipkow@15392: thus ?thesis by(auto simp add:image_def) nipkow@15392: qed nipkow@15392: qed nipkow@15392: next nipkow@15392: show "?r \ B" nipkow@15392: proof nipkow@15392: fix u assume "u \ ?r" nipkow@15392: then obtain i\<^isub>u where below: "i\<^isub>u < n" and nipkow@15392: or: "b = h i\<^isub>u \ u = h n \ h i\<^isub>u \ b \ h i\<^isub>u = u" nipkow@15392: by(auto simp:image_def) nipkow@15392: from or show "u \ B" nipkow@15392: proof nipkow@15392: assume [simp]: "b = h i\<^isub>u \ u = h n" nipkow@15392: have "u \ A" using card by auto nipkow@15392: moreover have "u \ b" using new below by auto nipkow@15392: ultimately show "u \ B" using A1 by blast nipkow@15392: next nipkow@15392: assume "h i\<^isub>u \ b \ h i\<^isub>u = u" nipkow@15392: moreover hence "u \ A" using card below by auto nipkow@15392: ultimately show "u \ B" using A1 by blast nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: show ?thesis nipkow@15392: proof cases nipkow@15392: assume "b = c" nipkow@15392: then moreover have "B = C" using A1 A2 notinB notinC by auto nipkow@15392: ultimately show ?thesis using IH[OF less] y z x x' by auto nipkow@15392: next nipkow@15392: assume diff: "b \ c" nipkow@15392: let ?D = "B - {c}" nipkow@15392: have B: "B = insert c ?D" and C: "C = insert b ?D" nipkow@15392: using A1 A2 notinB notinC diff by(blast elim!:equalityE)+ nipkow@15402: have "finite A" by(rule foldSet_imp_finite[OF Afoldx]) nipkow@15402: with A1 have "finite ?D" by simp nipkow@15392: then obtain d where Dfoldd: "(?D,d) \ foldSet f g e" nipkow@15392: using finite_imp_foldSet by rules nipkow@15392: moreover have cinB: "c \ B" using B by(auto) nipkow@15392: ultimately have "(B,g c \ d) \ foldSet f g e" nipkow@15392: by(rule Diff1_foldSet) nipkow@15392: hence "g c \ d = y" by(rule IH[OF less y]) nipkow@15392: moreover have "g b \ d = z" nipkow@15392: proof (rule IH[OF _ z]) nipkow@15392: let ?h = "%i. if h i = c then h n else h i" nipkow@15392: show "C = ?h`{i. i ?r" nipkow@15392: proof nipkow@15392: fix u assume "u \ C" nipkow@15392: hence uinA: "u \ A" and unotc: "u \ c" nipkow@15392: using A2 notinC by blast+ nipkow@15392: then obtain i\<^isub>u where below: "i\<^isub>u < Suc n" and [simp]: "u = h i\<^isub>u" nipkow@15392: using card by(auto simp:image_def) nipkow@15392: show "u \ ?r" nipkow@15392: proof cases nipkow@15392: assume "i\<^isub>u < n" nipkow@15392: thus ?thesis using unotc by(fastsimp) nipkow@15392: next nipkow@15392: assume "\ i\<^isub>u < n" nipkow@15392: with below have [simp]: "i\<^isub>u = n" by arith nipkow@15392: obtain i\<^isub>k where i\<^isub>k: "i\<^isub>k < Suc n" and [simp]: "c = h i\<^isub>k" nipkow@15392: using A2 card by blast nipkow@15392: have "i\<^isub>k < n" nipkow@15392: proof (rule ccontr) nipkow@15392: assume "\ i\<^isub>k < n" nipkow@15392: hence "i\<^isub>k = n" using i\<^isub>k by arith nipkow@15392: thus False using unotc by simp nipkow@15392: qed nipkow@15392: thus ?thesis by(auto simp add:image_def) nipkow@15392: qed nipkow@15392: qed nipkow@15392: next nipkow@15392: show "?r \ C" nipkow@15392: proof nipkow@15392: fix u assume "u \ ?r" nipkow@15392: then obtain i\<^isub>u where below: "i\<^isub>u < n" and nipkow@15392: or: "c = h i\<^isub>u \ u = h n \ h i\<^isub>u \ c \ h i\<^isub>u = u" nipkow@15392: by(auto simp:image_def) nipkow@15392: from or show "u \ C" nipkow@15392: proof nipkow@15392: assume [simp]: "c = h i\<^isub>u \ u = h n" nipkow@15392: have "u \ A" using card by auto nipkow@15392: moreover have "u \ c" using new below by auto nipkow@15392: ultimately show "u \ C" using A2 by blast nipkow@15392: next nipkow@15392: assume "h i\<^isub>u \ c \ h i\<^isub>u = u" nipkow@15392: moreover hence "u \ A" using card below by auto nipkow@15392: ultimately show "u \ C" using A2 by blast nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: next nipkow@15392: show "(C,g b \ d) \ foldSet f g e" using C notinB Dfoldd nipkow@15392: by fastsimp nipkow@15392: qed nipkow@15392: ultimately show ?thesis using x x' by(auto simp:AC) nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: nipkow@15392: (* The same proof, but using card nipkow@15392: lemma (in ACf) foldSet_determ_aux: nipkow@15392: "!!A x x'. \ card A < n; (A,x) : foldSet f g e; (A,x') : foldSet f g e \ nipkow@15392: \ x' = x" nipkow@15392: proof (induct n) nipkow@15392: case 0 thus ?case by simp nipkow@15392: next nipkow@15392: case (Suc n) nipkow@15392: have IH: "!!A x x'. \card A < n; (A,x) \ foldSet f g e; (A,x') \ foldSet f g e\ nipkow@15392: \ x' = x" and card: "card A < Suc n" nipkow@15392: and Afoldx: "(A, x) \ foldSet f g e" and Afoldy: "(A,x') \ foldSet f g e" . nipkow@15392: from card have "card A < n \ card A = n" by arith nipkow@15392: thus ?case nipkow@15392: proof nipkow@15392: assume less: "card A < n" nipkow@15392: show ?thesis by(rule IH[OF less Afoldx Afoldy]) nipkow@15392: next nipkow@15392: assume cardA: "card A = n" nipkow@15392: show ?thesis nipkow@15392: proof (rule foldSet.cases[OF Afoldx]) nipkow@15392: assume "(A, x) = ({}, e)" nipkow@15392: thus "x' = x" using Afoldy by (auto) nipkow@15392: next nipkow@15392: fix B b y nipkow@15392: assume eq1: "(A, x) = (insert b B, g b \ y)" nipkow@15392: and y: "(B,y) \ foldSet f g e" and notinB: "b \ B" nipkow@15392: hence A1: "A = insert b B" and x: "x = g b \ y" by auto nipkow@15392: show ?thesis nipkow@15392: proof (rule foldSet.cases[OF Afoldy]) nipkow@15392: assume "(A,x') = ({}, e)" nipkow@15392: thus ?thesis using A1 by auto nipkow@15392: next nipkow@15392: fix C c z nipkow@15392: assume eq2: "(A,x') = (insert c C, g c \ z)" nipkow@15392: and z: "(C,z) \ foldSet f g e" and notinC: "c \ C" nipkow@15392: hence A2: "A = insert c C" and x': "x' = g c \ z" by auto nipkow@15392: have finA: "finite A" by(rule foldSet_imp_finite[OF Afoldx]) nipkow@15392: with cardA A1 notinB have less: "card B < n" by simp nipkow@15392: show ?thesis nipkow@15392: proof cases nipkow@15392: assume "b = c" nipkow@15392: then moreover have "B = C" using A1 A2 notinB notinC by auto nipkow@15392: ultimately show ?thesis using IH[OF less] y z x x' by auto nipkow@15392: next nipkow@15392: assume diff: "b \ c" nipkow@15392: let ?D = "B - {c}" nipkow@15392: have B: "B = insert c ?D" and C: "C = insert b ?D" nipkow@15392: using A1 A2 notinB notinC diff by(blast elim!:equalityE)+ nipkow@15392: have "finite ?D" using finA A1 by simp nipkow@15392: then obtain d where Dfoldd: "(?D,d) \ foldSet f g e" nipkow@15392: using finite_imp_foldSet by rules nipkow@15392: moreover have cinB: "c \ B" using B by(auto) nipkow@15392: ultimately have "(B,g c \ d) \ foldSet f g e" nipkow@15392: by(rule Diff1_foldSet) nipkow@15392: hence "g c \ d = y" by(rule IH[OF less y]) nipkow@15392: moreover have "g b \ d = z" nipkow@15392: proof (rule IH[OF _ z]) nipkow@15392: show "card C < n" using C cardA A1 notinB finA cinB nipkow@15392: by(auto simp:card_Diff1_less) nipkow@15392: next nipkow@15392: show "(C,g b \ d) \ foldSet f g e" using C notinB Dfoldd nipkow@15392: by fastsimp nipkow@15392: qed nipkow@15392: ultimately show ?thesis using x x' by(auto simp:AC) nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: *) nipkow@15392: nipkow@15392: lemma (in ACf) foldSet_determ: nipkow@15392: "(A, x) : foldSet f g e ==> (A, y) : foldSet f g e ==> y = x" nipkow@15392: apply(frule foldSet_imp_finite) nipkow@15392: apply(simp add:finite_conv_nat_seg_image) nipkow@15392: apply(blast intro: foldSet_determ_aux [rule_format]) nipkow@15392: done nipkow@15392: nipkow@15392: lemma (in ACf) fold_equality: "(A, y) : foldSet f g e ==> fold f g e A = y" nipkow@15392: by (unfold fold_def) (blast intro: foldSet_determ) nipkow@15392: nipkow@15392: text{* The base case for @{text fold}: *} nipkow@15392: nipkow@15392: lemma fold_empty [simp]: "fold f g e {} = e" nipkow@15392: by (unfold fold_def) blast nipkow@15392: nipkow@15392: lemma (in ACf) fold_insert_aux: "x \ A ==> nipkow@15392: ((insert x A, v) : foldSet f g e) = nipkow@15392: (EX y. (A, y) : foldSet f g e & v = f (g x) y)" nipkow@15392: apply auto nipkow@15392: apply (rule_tac A1 = A and f1 = f in finite_imp_foldSet [THEN exE]) nipkow@15392: apply (fastsimp dest: foldSet_imp_finite) nipkow@15392: apply (blast intro: foldSet_determ) nipkow@15392: done nipkow@15392: nipkow@15392: text{* The recursion equation for @{text fold}: *} nipkow@15392: nipkow@15392: lemma (in ACf) fold_insert[simp]: nipkow@15392: "finite A ==> x \ A ==> fold f g e (insert x A) = f (g x) (fold f g e A)" nipkow@15392: apply (unfold fold_def) nipkow@15392: apply (simp add: fold_insert_aux) nipkow@15392: apply (rule the_equality) nipkow@15392: apply (auto intro: finite_imp_foldSet nipkow@15392: cong add: conj_cong simp add: fold_def [symmetric] fold_equality) nipkow@15392: done nipkow@15392: nipkow@15392: declare nipkow@15392: empty_foldSetE [rule del] foldSet.intros [rule del] nipkow@15392: -- {* Delete rules to do with @{text foldSet} relation. *} nipkow@15392: nipkow@15392: subsubsection{*Lemmas about @{text fold}*} nipkow@15392: nipkow@15392: lemma (in ACf) fold_commute: nipkow@15392: "finite A ==> (!!e. f (g x) (fold f g e A) = fold f g (f (g x) e) A)" nipkow@15392: apply (induct set: Finites, simp) nipkow@15392: apply (simp add: left_commute) nipkow@15392: done nipkow@15392: nipkow@15392: lemma (in ACf) fold_nest_Un_Int: nipkow@15392: "finite A ==> finite B nipkow@15392: ==> fold f g (fold f g e B) A = fold f g (fold f g e (A Int B)) (A Un B)" nipkow@15392: apply (induct set: Finites, simp) nipkow@15392: apply (simp add: fold_commute Int_insert_left insert_absorb) nipkow@15392: done nipkow@15392: nipkow@15392: lemma (in ACf) fold_nest_Un_disjoint: nipkow@15392: "finite A ==> finite B ==> A Int B = {} nipkow@15392: ==> fold f g e (A Un B) = fold f g (fold f g e B) A" nipkow@15392: by (simp add: fold_nest_Un_Int) nipkow@15392: nipkow@15392: lemma (in ACf) fold_reindex: nipkow@15392: assumes fin: "finite B" nipkow@15392: shows "inj_on h B \ fold f g e (h ` B) = fold f (g \ h) e B" nipkow@15392: using fin apply (induct) nipkow@15392: apply simp nipkow@15392: apply simp nipkow@15392: done nipkow@15392: nipkow@15392: lemma (in ACe) fold_Un_Int: nipkow@15392: "finite A ==> finite B ==> nipkow@15392: fold f g e A \ fold f g e B = nipkow@15392: fold f g e (A Un B) \ fold f g e (A Int B)" nipkow@15392: apply (induct set: Finites, simp) nipkow@15392: apply (simp add: AC insert_absorb Int_insert_left) nipkow@15392: done nipkow@15392: nipkow@15392: corollary (in ACe) fold_Un_disjoint: nipkow@15392: "finite A ==> finite B ==> A Int B = {} ==> nipkow@15392: fold f g e (A Un B) = fold f g e A \ fold f g e B" nipkow@15392: by (simp add: fold_Un_Int) nipkow@15392: nipkow@15392: lemma (in ACe) fold_UN_disjoint: nipkow@15392: "\ finite I; ALL i:I. finite (A i); nipkow@15392: ALL i:I. ALL j:I. i \ j --> A i Int A j = {} \ nipkow@15392: \ fold f g e (UNION I A) = nipkow@15392: fold f (%i. fold f g e (A i)) e I" nipkow@15392: apply (induct set: Finites, simp, atomize) nipkow@15392: apply (subgoal_tac "ALL i:F. x \ i") nipkow@15392: prefer 2 apply blast nipkow@15392: apply (subgoal_tac "A x Int UNION F A = {}") nipkow@15392: prefer 2 apply blast nipkow@15392: apply (simp add: fold_Un_disjoint) nipkow@15392: done nipkow@15392: nipkow@15392: lemma (in ACf) fold_cong: nipkow@15392: "finite A \ (!!x. x:A ==> g x = h x) ==> fold f g a A = fold f h a A" nipkow@15392: apply (subgoal_tac "ALL C. C <= A --> (ALL x:C. g x = h x) --> fold f g a C = fold f h a C") nipkow@15392: apply simp nipkow@15392: apply (erule finite_induct, simp) nipkow@15392: apply (simp add: subset_insert_iff, clarify) nipkow@15392: apply (subgoal_tac "finite C") nipkow@15392: prefer 2 apply (blast dest: finite_subset [COMP swap_prems_rl]) nipkow@15392: apply (subgoal_tac "C = insert x (C - {x})") nipkow@15392: prefer 2 apply blast nipkow@15392: apply (erule ssubst) nipkow@15392: apply (drule spec) nipkow@15392: apply (erule (1) notE impE) nipkow@15392: apply (simp add: Ball_def del: insert_Diff_single) nipkow@15392: done nipkow@15392: nipkow@15392: lemma (in ACe) fold_Sigma: "finite A ==> ALL x:A. finite (B x) ==> nipkow@15392: fold f (%x. fold f (g x) e (B x)) e A = nipkow@15392: fold f (split g) e (SIGMA x:A. B x)" nipkow@15392: apply (subst Sigma_def) nipkow@15392: apply (subst fold_UN_disjoint) nipkow@15392: apply assumption nipkow@15392: apply simp nipkow@15392: apply blast nipkow@15392: apply (erule fold_cong) nipkow@15392: apply (subst fold_UN_disjoint) nipkow@15392: apply simp nipkow@15392: apply simp nipkow@15392: apply blast nipkow@15392: apply (simp) nipkow@15392: done nipkow@15392: nipkow@15392: lemma (in ACe) fold_distrib: "finite A \ nipkow@15392: fold f (%x. f (g x) (h x)) e A = f (fold f g e A) (fold f h e A)" nipkow@15392: apply (erule finite_induct) nipkow@15392: apply simp nipkow@15392: apply (simp add:AC) nipkow@15392: done nipkow@15392: nipkow@15392: nipkow@15402: subsection {* Generalized summation over a set *} nipkow@15402: nipkow@15402: constdefs nipkow@15402: setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add" nipkow@15402: "setsum f A == if finite A then fold (op +) f 0 A else 0" nipkow@15402: 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 nipkow@15402: "_setsum" :: "idt => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_. _)" [0, 51, 10] 10) nipkow@15402: syntax (xsymbols) nipkow@15402: "_setsum" :: "idt => 'a set => 'b => 'b::comm_monoid_add" ("(3\_\_. _)" [0, 51, 10] 10) nipkow@15402: syntax (HTML output) nipkow@15402: "_setsum" :: "idt => 'a set => 'b => 'b::comm_monoid_add" ("(3\_\_. _)" [0, 51, 10] 10) nipkow@15402: nipkow@15402: translations -- {* Beware of argument permutation! *} nipkow@15402: "SUM i:A. b" == "setsum (%i. b) A" nipkow@15402: "\i\A. b" == "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 nipkow@15402: "_qsetsum" :: "idt \ bool \ 'a \ 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10) nipkow@15402: syntax (xsymbols) nipkow@15402: "_qsetsum" :: "idt \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) nipkow@15402: syntax (HTML output) nipkow@15402: "_qsetsum" :: "idt \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) nipkow@15402: nipkow@15402: translations nipkow@15402: "SUM x|P. t" => "setsum (%x. t) {x. P}" nipkow@15402: "\x|P. t" => "setsum (%x. t) {x. P}" nipkow@15402: nipkow@15402: text{* Finally we abbreviate @{term"\x\A. x"} by @{text"\A"}. *} nipkow@15402: nipkow@15402: syntax nipkow@15402: "_Setsum" :: "'a set => 'a::comm_monoid_mult" ("\_" [1000] 999) nipkow@15402: nipkow@15402: parse_translation {* nipkow@15402: let nipkow@15402: fun Setsum_tr [A] = Syntax.const "setsum" $ Abs ("", dummyT, Bound 0) $ A nipkow@15402: in [("_Setsum", Setsum_tr)] end; nipkow@15402: *} nipkow@15402: nipkow@15402: print_translation {* nipkow@15402: let nipkow@15402: fun setsum_tr' [Abs(_,_,Bound 0), A] = Syntax.const "_Setsum" $ A nipkow@15402: | setsum_tr' [Abs(x,Tx,t), Const ("Collect",_) $ Abs(y,Ty,P)] = nipkow@15402: if x<>y then raise Match nipkow@15402: else let val x' = Syntax.mark_bound x nipkow@15402: val t' = subst_bound(x',t) nipkow@15402: val P' = subst_bound(x',P) nipkow@15402: in Syntax.const "_qsetsum" $ Syntax.mark_bound x $ P' $ t' end nipkow@15402: in nipkow@15402: [("setsum", setsum_tr')] nipkow@15402: end nipkow@15402: *} nipkow@15402: nipkow@15402: lemma setsum_empty [simp]: "setsum f {} = 0" nipkow@15402: by (simp add: setsum_def) nipkow@15402: nipkow@15402: lemma setsum_insert [simp]: nipkow@15402: "finite F ==> a \ F ==> setsum f (insert a F) = f a + setsum f F" nipkow@15402: by (simp add: setsum_def ACf.fold_insert [OF ACf_add]) nipkow@15402: nipkow@15402: lemma setsum_reindex: nipkow@15402: "inj_on f B ==> setsum h (f ` B) = setsum (h \ f) B" nipkow@15402: by(auto simp add: setsum_def ACf.fold_reindex[OF ACf_add] dest!:finite_imageD) nipkow@15402: nipkow@15402: lemma setsum_reindex_id: nipkow@15402: "inj_on f B ==> setsum f B = setsum id (f ` B)" nipkow@15402: by (auto simp add: setsum_reindex) nipkow@15402: nipkow@15402: lemma setsum_cong: nipkow@15402: "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B" nipkow@15402: by(fastsimp simp: setsum_def intro: ACf.fold_cong[OF ACf_add]) nipkow@15402: nipkow@15402: lemma setsum_reindex_cong: nipkow@15402: "[|inj_on f A; B = f ` A; !!a. g a = h (f a)|] nipkow@15402: ==> setsum h B = setsum g A" nipkow@15402: by (simp add: setsum_reindex cong: setsum_cong) nipkow@15402: nipkow@15402: lemma setsum_0: "setsum (%i. 0) A = 0" nipkow@15402: apply (clarsimp simp: setsum_def) nipkow@15402: apply (erule finite_induct, auto simp:ACf.fold_insert [OF ACf_add]) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setsum_0': "ALL a:F. f a = 0 ==> setsum f F = 0" nipkow@15402: apply (subgoal_tac "setsum f F = setsum (%x. 0) F") nipkow@15402: apply (erule ssubst, rule setsum_0) nipkow@15402: apply (rule setsum_cong, auto) nipkow@15402: done nipkow@15402: nipkow@15402: 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@15402: by(simp add: setsum_def ACe.fold_Un_Int[OF ACe_add,symmetric]) nipkow@15402: nipkow@15402: 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@15402: by (subst setsum_Un_Int [symmetric], auto) nipkow@15402: nipkow@15402: (* FIXME get rid of finite I. If infinite, rhs is directly 0, and UNION I A nipkow@15402: is also infinite and hence also 0 *) nipkow@15402: lemma setsum_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: setsum f (UNION I A) = (\i\I. setsum f (A i))" nipkow@15402: by(simp add: setsum_def ACe.fold_UN_disjoint[OF ACe_add] cong: setsum_cong) nipkow@15402: nipkow@15402: nipkow@15402: (* FIXME get rid of finite C. If infinite, rhs is directly 0, and Union C nipkow@15402: is also infinite and hence also 0 *) nipkow@15402: lemma setsum_Union_disjoint: nipkow@15402: "finite C ==> (ALL A:C. finite A) ==> nipkow@15402: (ALL A:C. ALL B:C. A \ B --> A Int B = {}) ==> nipkow@15402: setsum f (Union C) = setsum (setsum f) C" nipkow@15402: apply (frule setsum_UN_disjoint [of C id f]) nipkow@15402: apply (unfold Union_def id_def, assumption+) nipkow@15402: done nipkow@15402: nipkow@15402: (* FIXME get rid of finite A. If infinite, lhs is directly 0, and SIGMA A B nipkow@15402: is either infinite or empty, and in both cases the rhs is also 0 *) nipkow@15402: lemma setsum_Sigma: "finite A ==> ALL x:A. finite (B x) ==> nipkow@15402: (\x\A. (\y\B x. f x y)) = nipkow@15402: (\z\(SIGMA x:A. B x). f (fst z) (snd z))" nipkow@15402: by(simp add:setsum_def ACe.fold_Sigma[OF ACe_add] split_def cong:setsum_cong) nipkow@15402: nipkow@15402: lemma setsum_cartesian_product: "finite A ==> finite B ==> nipkow@15402: (\x\A. (\y\B. f x y)) = nipkow@15402: (\z\A <*> B. f (fst z) (snd z))" nipkow@15402: by (erule setsum_Sigma, auto) nipkow@15402: nipkow@15402: lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)" nipkow@15402: by(simp add:setsum_def ACe.fold_distrib[OF ACe_add]) nipkow@15402: 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@15402: apply (case_tac "finite A") nipkow@15402: prefer 2 apply (simp add: setsum_def) nipkow@15402: apply (erule rev_mp) nipkow@15402: apply (erule finite_induct, auto) nipkow@15402: 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@15402: by (induct set: Finites) auto nipkow@15402: nipkow@15402: lemma setsum_Un_nat: "finite A ==> finite B ==> nipkow@15402: (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@15402: by (subst setsum_Un_Int [symmetric], auto simp add: ring_eq_simps) nipkow@15402: nipkow@15402: lemma setsum_Un: "finite A ==> finite B ==> nipkow@15402: (setsum f (A Un B) :: 'a :: ab_group_add) = nipkow@15402: setsum f A + setsum f B - setsum f (A Int B)" nipkow@15402: by (subst setsum_Un_Int [symmetric], auto simp add: ring_eq_simps) nipkow@15402: nipkow@15402: lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) = nipkow@15402: (if a:A then setsum f A - f a else setsum f A)" nipkow@15402: apply (case_tac "finite A") nipkow@15402: prefer 2 apply (simp add: setsum_def) nipkow@15402: apply (erule finite_induct) nipkow@15402: apply (auto simp add: insert_Diff_if) nipkow@15402: apply (drule_tac a = a in mk_disjoint_insert, auto) nipkow@15402: 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@15402: by (erule finite_induct) (auto simp add: insert_Diff_if) nipkow@15402: nipkow@15402: (* By Jeremy Siek: *) nipkow@15402: nipkow@15402: lemma setsum_diff_nat: nipkow@15402: assumes finB: "finite B" nipkow@15402: shows "B \ A \ (setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)" nipkow@15402: using finB nipkow@15402: 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 nipkow@15402: proof (induct) nipkow@15402: case empty nipkow@15402: thus ?case by auto nipkow@15402: next nipkow@15402: case (insert x F) nipkow@15402: thus ?case using le finiteB nipkow@15402: by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb) nipkow@15402: qed nipkow@15402: qed nipkow@15402: nipkow@15402: lemma setsum_mono: nipkow@15402: assumes le: "\i. i\K \ f (i::'a) \ ((g i)::('b::{comm_monoid_add, pordered_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 nipkow@15402: proof (induct) nipkow@15402: case empty nipkow@15402: thus ?case by simp nipkow@15402: next nipkow@15402: case insert nipkow@15402: thus ?case using add_mono nipkow@15402: by force 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@15402: lemma setsum_mono2_nat: nipkow@15402: assumes fin: "finite B" and sub: "A \ B" nipkow@15402: shows "setsum f A \ (setsum f B :: nat)" nipkow@15402: proof - nipkow@15402: have "setsum f A \ setsum f A + setsum f (B-A)" by arith nipkow@15402: also have "\ = setsum f (A \ (B-A))" using fin finite_subset[OF sub fin] nipkow@15402: by (simp add:setsum_Un_disjoint del:Un_Diff_cancel) nipkow@15402: also have "A \ (B-A) = B" using sub by blast nipkow@15402: finally show ?thesis . nipkow@15402: qed nipkow@15402: nipkow@15402: lemma setsum_negf: "finite A ==> setsum (%x. - (f x)::'a::ab_group_add) A = nipkow@15402: - setsum f A" nipkow@15402: by (induct set: Finites, auto) nipkow@15402: nipkow@15402: lemma setsum_subtractf: "finite A ==> setsum (%x. ((f x)::'a::ab_group_add) - g x) A = nipkow@15402: setsum f A - setsum g A" nipkow@15402: by (simp add: diff_minus setsum_addf setsum_negf) nipkow@15402: nipkow@15402: lemma setsum_nonneg: "[| finite A; nipkow@15402: \x \ A. (0::'a::{pordered_ab_semigroup_add, comm_monoid_add}) \ f x |] ==> nipkow@15402: 0 \ setsum f A"; nipkow@15402: apply (induct set: Finites, auto) nipkow@15402: apply (subgoal_tac "0 + 0 \ f x + setsum f F", simp) nipkow@15402: apply (blast intro: add_mono) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setsum_nonpos: "[| finite A; nipkow@15402: \x \ A. f x \ (0::'a::{pordered_ab_semigroup_add, comm_monoid_add}) |] ==> nipkow@15402: setsum f A \ 0"; nipkow@15402: apply (induct set: Finites, auto) nipkow@15402: apply (subgoal_tac "f x + setsum f F \ 0 + 0", simp) nipkow@15402: apply (blast intro: add_mono) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setsum_mult: nipkow@15402: fixes f :: "'a => ('b::semiring_0_cancel)" 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 nipkow@15402: 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: nipkow@15402: lemma setsum_abs: nipkow@15402: fixes f :: "'a => ('b::lordered_ab_group_abs)" nipkow@15402: assumes fin: "finite A" nipkow@15402: shows "abs (setsum f A) \ setsum (%i. abs(f i)) A" nipkow@15402: using fin nipkow@15402: proof (induct) nipkow@15402: case empty thus ?case by simp nipkow@15402: next nipkow@15402: case (insert x A) nipkow@15402: thus ?case by (auto intro: abs_triangle_ineq order_trans) nipkow@15402: qed nipkow@15402: nipkow@15402: lemma setsum_abs_ge_zero: nipkow@15402: fixes f :: "'a => ('b::lordered_ab_group_abs)" nipkow@15402: assumes fin: "finite A" nipkow@15402: shows "0 \ setsum (%i. abs(f i)) A" nipkow@15402: using fin nipkow@15402: proof (induct) nipkow@15402: case empty thus ?case by simp nipkow@15402: next nipkow@15402: case (insert x A) thus ?case by (auto intro: order_trans) nipkow@15402: qed nipkow@15402: nipkow@15402: nipkow@15402: subsection {* Generalized product over a set *} nipkow@15402: nipkow@15402: constdefs nipkow@15402: setprod :: "('a => 'b) => 'a set => 'b::comm_monoid_mult" nipkow@15402: "setprod f A == if finite A then fold (op *) f 1 A else 1" nipkow@15402: nipkow@15402: syntax nipkow@15402: "_setprod" :: "idt => 'a set => 'b => 'b::comm_monoid_mult" ("(3\_:_. _)" [0, 51, 10] 10) nipkow@15402: nipkow@15402: syntax (xsymbols) nipkow@15402: "_setprod" :: "idt => 'a set => 'b => 'b::comm_monoid_mult" ("(3\_\_. _)" [0, 51, 10] 10) nipkow@15402: syntax (HTML output) nipkow@15402: "_setprod" :: "idt => 'a set => 'b => 'b::comm_monoid_mult" ("(3\_\_. _)" [0, 51, 10] 10) nipkow@15402: translations nipkow@15402: "\i:A. b" == "setprod (%i. b) A" -- {* Beware of argument permutation! *} nipkow@15402: nipkow@15402: syntax nipkow@15402: "_Setprod" :: "'a set => 'a::comm_monoid_mult" ("\_" [1000] 999) nipkow@15402: nipkow@15402: parse_translation {* nipkow@15402: let nipkow@15402: fun Setprod_tr [A] = Syntax.const "setprod" $ Abs ("", dummyT, Bound 0) $ A nipkow@15402: in [("_Setprod", Setprod_tr)] end; nipkow@15402: *} nipkow@15402: print_translation {* nipkow@15402: let fun setprod_tr' [Abs(x,Tx,t), A] = nipkow@15402: if t = Bound 0 then Syntax.const "_Setprod" $ A else raise Match nipkow@15402: in nipkow@15402: [("setprod", setprod_tr')] nipkow@15402: end nipkow@15402: *} nipkow@15402: nipkow@15402: nipkow@15402: lemma setprod_empty [simp]: "setprod f {} = 1" nipkow@15402: by (auto simp add: setprod_def) nipkow@15402: nipkow@15402: lemma setprod_insert [simp]: "[| finite A; a \ A |] ==> nipkow@15402: setprod f (insert a A) = f a * setprod f A" nipkow@15402: by (simp add: setprod_def ACf.fold_insert [OF ACf_mult]) nipkow@15402: nipkow@15402: lemma setprod_reindex: nipkow@15402: "inj_on f B ==> setprod h (f ` B) = setprod (h \ f) B" nipkow@15402: by(auto simp: setprod_def ACf.fold_reindex[OF ACf_mult] 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@15402: by(fastsimp simp: setprod_def intro: ACf.fold_cong[OF ACf_mult]) nipkow@15402: 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@15402: by (frule setprod_reindex, simp) nipkow@15402: nipkow@15402: nipkow@15402: lemma setprod_1: "setprod (%i. 1) A = 1" nipkow@15402: apply (case_tac "finite A") nipkow@15402: apply (erule finite_induct, auto simp add: mult_ac) nipkow@15402: apply (simp add: setprod_def) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_1': "ALL a:F. f a = 1 ==> setprod f F = 1" nipkow@15402: apply (subgoal_tac "setprod f F = setprod (%x. 1) F") nipkow@15402: apply (erule ssubst, rule setprod_1) nipkow@15402: apply (rule setprod_cong, auto) nipkow@15402: done 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@15402: by(simp add: setprod_def ACe.fold_Un_Int[OF ACe_mult,symmetric]) 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@15402: by (subst setprod_Un_Int [symmetric], auto) nipkow@15402: 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" nipkow@15402: by(simp add: setprod_def ACe.fold_UN_disjoint[OF ACe_mult] cong: setprod_cong) nipkow@15402: nipkow@15402: lemma setprod_Union_disjoint: nipkow@15402: "finite C ==> (ALL A:C. finite A) ==> nipkow@15402: (ALL A:C. ALL B:C. A \ B --> A Int B = {}) ==> nipkow@15402: setprod f (Union C) = setprod (setprod f) C" nipkow@15402: apply (frule setprod_UN_disjoint [of C id f]) nipkow@15402: apply (unfold Union_def id_def, assumption+) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==> nipkow@15402: (\x:A. (\y: B x. f x y)) = nipkow@15402: (\z:(SIGMA x:A. B x). f (fst z) (snd z))" nipkow@15402: by(simp add:setprod_def ACe.fold_Sigma[OF ACe_mult] split_def cong:setprod_cong) nipkow@15402: nipkow@15402: lemma setprod_cartesian_product: "finite A ==> finite B ==> nipkow@15402: (\x:A. (\y: B. f x y)) = nipkow@15402: (\z:(A <*> B). f (fst z) (snd z))" nipkow@15402: by (erule setprod_Sigma, auto) nipkow@15402: nipkow@15402: lemma setprod_timesf: nipkow@15402: "setprod (%x. f x * g x) A = (setprod f A * setprod g A)" nipkow@15402: by(simp add:setprod_def ACe.fold_distrib[OF ACe_mult]) 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@15402: "finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))" nipkow@15402: by (induct set: Finites) auto nipkow@15402: nipkow@15402: lemma setprod_zero: nipkow@15402: "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1_cancel) ==> setprod f A = 0" nipkow@15402: apply (induct set: Finites, force, clarsimp) nipkow@15402: apply (erule disjE, auto) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_nonneg [rule_format]: nipkow@15402: "(ALL x: A. (0::'a::ordered_idom) \ f x) --> 0 \ setprod f A" nipkow@15402: apply (case_tac "finite A") nipkow@15402: apply (induct set: Finites, force, clarsimp) nipkow@15402: apply (subgoal_tac "0 * 0 \ f x * setprod f F", force) nipkow@15402: apply (rule mult_mono, assumption+) nipkow@15402: apply (auto simp add: setprod_def) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::ordered_idom) < f x) nipkow@15402: --> 0 < setprod f A" nipkow@15402: apply (case_tac "finite A") nipkow@15402: apply (induct set: Finites, force, clarsimp) nipkow@15402: apply (subgoal_tac "0 * 0 < f x * setprod f F", force) nipkow@15402: apply (rule mult_strict_mono, assumption+) nipkow@15402: apply (auto simp add: setprod_def) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_nonzero [rule_format]: nipkow@15402: "(ALL x y. (x::'a::comm_semiring_1_cancel) * y = 0 --> x = 0 | y = 0) ==> nipkow@15402: finite A ==> (ALL x: A. f x \ (0::'a)) --> setprod f A \ 0" nipkow@15402: apply (erule finite_induct, auto) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_zero_eq: nipkow@15402: "(ALL x y. (x::'a::comm_semiring_1_cancel) * y = 0 --> x = 0 | y = 0) ==> nipkow@15402: finite A ==> (setprod f A = (0::'a)) = (EX x: A. f x = 0)" nipkow@15402: apply (insert setprod_zero [of A f] setprod_nonzero [of A f], blast) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_nonzero_field: nipkow@15402: "finite A ==> (ALL x: A. f x \ (0::'a::field)) ==> setprod f A \ 0" nipkow@15402: apply (rule setprod_nonzero, auto) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_zero_eq_field: nipkow@15402: "finite A ==> (setprod f A = (0::'a::field)) = (EX x: A. f x = 0)" nipkow@15402: apply (rule setprod_zero_eq, auto) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \ 0) ==> nipkow@15402: (setprod f (A Un B) :: 'a ::{field}) nipkow@15402: = setprod f A * setprod f B / setprod f (A Int B)" nipkow@15402: apply (subst setprod_Un_Int [symmetric], auto) nipkow@15402: apply (subgoal_tac "finite (A Int B)") nipkow@15402: apply (frule setprod_nonzero_field [of "A Int B" f], assumption) nipkow@15402: apply (subst times_divide_eq_right [THEN sym], auto simp add: divide_self) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_diff1: "finite A ==> f a \ 0 ==> nipkow@15402: (setprod f (A - {a}) :: 'a :: {field}) = nipkow@15402: (if a:A then setprod f A / f a else setprod f A)" nipkow@15402: apply (erule finite_induct) nipkow@15402: apply (auto simp add: insert_Diff_if) nipkow@15402: apply (subgoal_tac "f a * setprod f F / f a = setprod f F * f a / f a") nipkow@15402: apply (erule ssubst) nipkow@15402: apply (subst times_divide_eq_right [THEN sym]) nipkow@15402: apply (auto simp add: mult_ac times_divide_eq_right divide_self) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_inversef: "finite A ==> nipkow@15402: ALL x: A. f x \ (0::'a::{field,division_by_zero}) ==> nipkow@15402: setprod (inverse \ f) A = inverse (setprod f A)" nipkow@15402: apply (erule finite_induct) nipkow@15402: apply (simp, simp) nipkow@15402: done nipkow@15402: nipkow@15402: lemma setprod_dividef: nipkow@15402: "[|finite A; nipkow@15402: \x \ A. g x \ (0::'a::{field,division_by_zero})|] nipkow@15402: ==> setprod (%x. f x / g x) A = setprod f A / setprod g A" nipkow@15402: apply (subgoal_tac nipkow@15402: "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \ g) x) A") nipkow@15402: apply (erule ssubst) nipkow@15402: apply (subst divide_inverse) nipkow@15402: apply (subst setprod_timesf) nipkow@15402: apply (subst setprod_inversef, assumption+, rule refl) nipkow@15402: apply (rule setprod_cong, rule refl) nipkow@15402: apply (subst divide_inverse, auto) nipkow@15402: done nipkow@15402: wenzelm@12396: subsection {* Finite cardinality *} wenzelm@12396: nipkow@15402: text {* This definition, although traditional, is ugly to work with: nipkow@15402: @{text "card A == LEAST n. EX f. A = {f i | i. i < n}"}. nipkow@15402: But now that we have @{text setsum} things are easy: wenzelm@12396: *} wenzelm@12396: wenzelm@12396: constdefs wenzelm@12396: card :: "'a set => nat" nipkow@15402: "card A == setsum (%x. 1::nat) A" wenzelm@12396: wenzelm@12396: lemma card_empty [simp]: "card {} = 0" nipkow@15402: by (simp add: card_def) nipkow@15402: nipkow@15402: lemma card_eq_setsum: "card A = setsum (%x. 1) A" nipkow@15402: by (simp add: card_def) wenzelm@12396: wenzelm@12396: lemma card_insert_disjoint [simp]: wenzelm@12396: "finite A ==> x \ A ==> card (insert x A) = Suc(card A)" nipkow@15402: by(simp add: card_def ACf.fold_insert[OF ACf_add]) nipkow@15402: nipkow@15402: lemma card_insert_if: nipkow@15402: "finite A ==> card (insert x A) = (if x:A then card A else Suc(card(A)))" nipkow@15402: by (simp add: insert_absorb) wenzelm@12396: wenzelm@12396: lemma card_0_eq [simp]: "finite A ==> (card A = 0) = (A = {})" wenzelm@12396: apply auto paulson@14208: apply (drule_tac a = x in mk_disjoint_insert, clarify) nipkow@15402: apply (auto) wenzelm@12396: done wenzelm@12396: wenzelm@12396: lemma card_Suc_Diff1: "finite A ==> x: A ==> Suc (card (A - {x})) = card A" nipkow@14302: apply(rule_tac t = A in insert_Diff [THEN subst], assumption) nipkow@14302: apply(simp del:insert_Diff_single) nipkow@14302: done wenzelm@12396: wenzelm@12396: lemma card_Diff_singleton: wenzelm@12396: "finite A ==> x: A ==> card (A - {x}) = card A - 1" wenzelm@12396: by (simp add: card_Suc_Diff1 [symmetric]) wenzelm@12396: wenzelm@12396: lemma card_Diff_singleton_if: wenzelm@12396: "finite A ==> card (A-{x}) = (if x : A then card A - 1 else card A)" wenzelm@12396: by (simp add: card_Diff_singleton) wenzelm@12396: wenzelm@12396: lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))" wenzelm@12396: by (simp add: card_insert_if card_Suc_Diff1) wenzelm@12396: wenzelm@12396: lemma card_insert_le: "finite A ==> card A <= card (insert x A)" wenzelm@12396: by (simp add: card_insert_if) wenzelm@12396: nipkow@15402: lemma card_mono: "\ finite B; A \ B \ \ card A \ card B" nipkow@15402: by (simp add: card_def setsum_mono2_nat) nipkow@15402: wenzelm@12396: lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)" paulson@14208: apply (induct set: Finites, simp, clarify) wenzelm@12396: apply (subgoal_tac "finite A & A - {x} <= F") paulson@14208: prefer 2 apply (blast intro: finite_subset, atomize) wenzelm@12396: apply (drule_tac x = "A - {x}" in spec) wenzelm@12396: apply (simp add: card_Diff_singleton_if split add: split_if_asm) paulson@14208: apply (case_tac "card A", auto) wenzelm@12396: done wenzelm@12396: wenzelm@12396: lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B" wenzelm@12396: apply (simp add: psubset_def linorder_not_le [symmetric]) wenzelm@12396: apply (blast dest: card_seteq) wenzelm@12396: done wenzelm@12396: wenzelm@12396: lemma card_Un_Int: "finite A ==> finite B wenzelm@12396: ==> card A + card B = card (A Un B) + card (A Int B)" nipkow@15402: by(simp add:card_def setsum_Un_Int) wenzelm@12396: wenzelm@12396: lemma card_Un_disjoint: "finite A ==> finite B wenzelm@12396: ==> A Int B = {} ==> card (A Un B) = card A + card B" wenzelm@12396: by (simp add: card_Un_Int) wenzelm@12396: wenzelm@12396: lemma card_Diff_subset: nipkow@15402: "finite B ==> B <= A ==> card (A - B) = card A - card B" nipkow@15402: by(simp add:card_def setsum_diff_nat) wenzelm@12396: wenzelm@12396: lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A" wenzelm@12396: apply (rule Suc_less_SucD) wenzelm@12396: apply (simp add: card_Suc_Diff1) wenzelm@12396: done wenzelm@12396: wenzelm@12396: lemma card_Diff2_less: wenzelm@12396: "finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A" wenzelm@12396: apply (case_tac "x = y") wenzelm@12396: apply (simp add: card_Diff1_less) wenzelm@12396: apply (rule less_trans) wenzelm@12396: prefer 2 apply (auto intro!: card_Diff1_less) wenzelm@12396: done wenzelm@12396: wenzelm@12396: lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A" wenzelm@12396: apply (case_tac "x : A") wenzelm@12396: apply (simp_all add: card_Diff1_less less_imp_le) wenzelm@12396: done wenzelm@12396: wenzelm@12396: lemma card_psubset: "finite B ==> A \ B ==> card A < card B ==> A < B" paulson@14208: by (erule psubsetI, blast) wenzelm@12396: paulson@14889: lemma insert_partition: nipkow@15402: "\ x \ F; \c1 \ insert x F. \c2 \ insert x F. c1 \ c2 \ c1 \ c2 = {} \ nipkow@15402: \ x \ \ F = {}" paulson@14889: by auto paulson@14889: paulson@14889: (* main cardinality theorem *) paulson@14889: lemma card_partition [rule_format]: paulson@14889: "finite C ==> paulson@14889: finite (\ C) --> paulson@14889: (\c\C. card c = k) --> paulson@14889: (\c1 \ C. \c2 \ C. c1 \ c2 --> c1 \ c2 = {}) --> paulson@14889: k * card(C) = card (\ C)" paulson@14889: apply (erule finite_induct, simp) paulson@14889: apply (simp add: card_insert_disjoint card_Un_disjoint insert_partition paulson@14889: finite_subset [of _ "\ (insert x F)"]) paulson@14889: done paulson@14889: wenzelm@12396: nipkow@15402: lemma setsum_constant_nat: nipkow@15402: "finite A ==> (\x\A. y) = (card A) * y" nipkow@15402: -- {* Generalized to any @{text comm_semiring_1_cancel} in nipkow@15402: @{text IntDef} as @{text setsum_constant}. *} nipkow@15402: by (erule finite_induct, auto) nipkow@15402: nipkow@15402: lemma setprod_constant: "finite A ==> (\x: A. (y::'a::recpower)) = y^(card A)" nipkow@15402: apply (erule finite_induct) nipkow@15402: apply (auto simp add: power_Suc) nipkow@15402: done nipkow@15402: nipkow@15402: nipkow@15402: subsubsection {* Cardinality of unions *} nipkow@15402: nipkow@15402: lemma card_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: card (UNION I A) = (\i\I. card(A i))" nipkow@15402: apply (simp add: card_def) nipkow@15402: apply (subgoal_tac nipkow@15402: "setsum (%i. card (A i)) I = setsum (%i. (setsum (%x. 1) (A i))) I") nipkow@15402: apply (simp add: setsum_UN_disjoint) nipkow@15402: apply (simp add: setsum_constant_nat cong: setsum_cong) nipkow@15402: done nipkow@15402: nipkow@15402: lemma card_Union_disjoint: nipkow@15402: "finite C ==> (ALL A:C. finite A) ==> nipkow@15402: (ALL A:C. ALL B:C. A \ B --> A Int B = {}) ==> nipkow@15402: card (Union C) = setsum card C" nipkow@15402: apply (frule card_UN_disjoint [of C id]) nipkow@15402: apply (unfold Union_def id_def, assumption+) nipkow@15402: done nipkow@15402: wenzelm@12396: subsubsection {* Cardinality of image *} wenzelm@12396: wenzelm@12396: lemma card_image_le: "finite A ==> card (f ` A) <= card A" paulson@14208: apply (induct set: Finites, simp) wenzelm@12396: apply (simp add: le_SucI finite_imageI card_insert_if) wenzelm@12396: done wenzelm@12396: nipkow@15402: lemma card_image: "inj_on f A ==> card (f ` A) = card A" nipkow@15402: by(simp add:card_def setsum_reindex o_def) wenzelm@12396: wenzelm@12396: lemma endo_inj_surj: "finite A ==> f ` A \ A ==> inj_on f A ==> f ` A = A" wenzelm@12396: by (simp add: card_seteq card_image) wenzelm@12396: nipkow@15111: lemma eq_card_imp_inj_on: nipkow@15111: "[| finite A; card(f ` A) = card A |] ==> inj_on f A" nipkow@15111: apply(induct rule:finite_induct) nipkow@15111: apply simp nipkow@15111: apply(frule card_image_le[where f = f]) nipkow@15111: apply(simp add:card_insert_if split:if_splits) nipkow@15111: done nipkow@15111: nipkow@15111: lemma inj_on_iff_eq_card: nipkow@15111: "finite A ==> inj_on f A = (card(f ` A) = card A)" nipkow@15111: by(blast intro: card_image eq_card_imp_inj_on) nipkow@15111: wenzelm@12396: nipkow@15402: lemma card_inj_on_le: nipkow@15402: "[|inj_on f A; f ` A \ B; finite B |] ==> card A \ card B" nipkow@15402: apply (subgoal_tac "finite A") nipkow@15402: apply (force intro: card_mono simp add: card_image [symmetric]) nipkow@15402: apply (blast intro: finite_imageD dest: finite_subset) nipkow@15402: done nipkow@15402: nipkow@15402: lemma card_bij_eq: nipkow@15402: "[|inj_on f A; f ` A \ B; inj_on g B; g ` B \ A; nipkow@15402: finite A; finite B |] ==> card A = card B" nipkow@15402: by (auto intro: le_anti_sym card_inj_on_le) nipkow@15402: nipkow@15402: nipkow@15402: subsubsection {* Cardinality of products *} nipkow@15402: nipkow@15402: (* nipkow@15402: lemma SigmaI_insert: "y \ A ==> nipkow@15402: (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \ (SIGMA x: A. B x))" nipkow@15402: by auto nipkow@15402: *) nipkow@15402: nipkow@15402: lemma card_SigmaI [simp]: nipkow@15402: "\ finite A; ALL a:A. finite (B a) \ nipkow@15402: \ card (SIGMA x: A. B x) = (\a\A. card (B a))" nipkow@15402: by(simp add:card_def setsum_Sigma) nipkow@15402: nipkow@15402: (* FIXME get rid of prems *) nipkow@15402: lemma card_cartesian_product: nipkow@15402: "[| finite A; finite B |] ==> card (A <*> B) = card(A) * card(B)" nipkow@15402: by (simp add: setsum_constant_nat) nipkow@15402: nipkow@15402: (* FIXME should really be a consequence of card_cartesian_product *) nipkow@15402: lemma card_cartesian_product_singleton: "card({x} <*> A) = card(A)" nipkow@15402: apply (subgoal_tac "inj_on (%y .(x,y)) A") nipkow@15402: apply (frule card_image) nipkow@15402: apply (subgoal_tac "(Pair x ` A) = {x} <*> A") nipkow@15402: apply (auto simp add: inj_on_def) nipkow@15402: done nipkow@15402: nipkow@15402: wenzelm@12396: subsubsection {* Cardinality of the Powerset *} wenzelm@12396: wenzelm@12396: lemma card_Pow: "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A" (* FIXME numeral 2 (!?) *) wenzelm@12396: apply (induct set: Finites) wenzelm@12396: apply (simp_all add: Pow_insert) paulson@14208: apply (subst card_Un_disjoint, blast) paulson@14208: apply (blast intro: finite_imageI, blast) wenzelm@12396: apply (subgoal_tac "inj_on (insert x) (Pow F)") wenzelm@12396: apply (simp add: card_image Pow_insert) wenzelm@12396: apply (unfold inj_on_def) wenzelm@12396: apply (blast elim!: equalityE) wenzelm@12396: done wenzelm@12396: nipkow@15392: text {* Relates to equivalence classes. Based on a theorem of nipkow@15392: F. Kammüller's. *} wenzelm@12396: wenzelm@12396: lemma dvd_partition: nipkow@15392: "finite (Union C) ==> wenzelm@12396: ALL c : C. k dvd card c ==> paulson@14430: (ALL c1: C. ALL c2: C. c1 \ c2 --> c1 Int c2 = {}) ==> wenzelm@12396: k dvd card (Union C)" nipkow@15392: apply(frule finite_UnionD) nipkow@15392: apply(rotate_tac -1) paulson@14208: apply (induct set: Finites, simp_all, clarify) wenzelm@12396: apply (subst card_Un_disjoint) wenzelm@12396: apply (auto simp add: dvd_add disjoint_eq_subset_Compl) wenzelm@12396: done wenzelm@12396: wenzelm@12396: nipkow@15392: subsubsection {* Theorems about @{text "choose"} *} wenzelm@12396: wenzelm@12396: text {* nipkow@15392: \medskip Basic theorem about @{text "choose"}. By Florian nipkow@15392: Kamm\"uller, tidied by LCP. wenzelm@12396: *} wenzelm@12396: nipkow@15392: lemma card_s_0_eq_empty: nipkow@15392: "finite A ==> card {B. B \ A & card B = 0} = 1" nipkow@15392: apply (simp cong add: conj_cong add: finite_subset [THEN card_0_eq]) nipkow@15392: apply (simp cong add: rev_conj_cong) nipkow@15392: done wenzelm@12396: nipkow@15392: lemma choose_deconstruct: "finite M ==> x \ M nipkow@15392: ==> {s. s <= insert x M & card(s) = Suc k} nipkow@15392: = {s. s <= M & card(s) = Suc k} Un nipkow@15392: {s. EX t. t <= M & card(t) = k & s = insert x t}" nipkow@15392: apply safe nipkow@15392: apply (auto intro: finite_subset [THEN card_insert_disjoint]) nipkow@15392: apply (drule_tac x = "xa - {x}" in spec) nipkow@15392: apply (subgoal_tac "x \ xa", auto) nipkow@15392: apply (erule rev_mp, subst card_Diff_singleton) nipkow@15392: apply (auto intro: finite_subset) wenzelm@12396: done wenzelm@12396: nipkow@15392: text{*There are as many subsets of @{term A} having cardinality @{term k} nipkow@15392: as there are sets obtained from the former by inserting a fixed element nipkow@15392: @{term x} into each.*} nipkow@15392: lemma constr_bij: nipkow@15392: "[|finite A; x \ A|] ==> nipkow@15392: card {B. EX C. C <= A & card(C) = k & B = insert x C} = nipkow@15392: card {B. B <= A & card(B) = k}" nipkow@15392: apply (rule_tac f = "%s. s - {x}" and g = "insert x" in card_bij_eq) nipkow@15392: apply (auto elim!: equalityE simp add: inj_on_def) nipkow@15392: apply (subst Diff_insert0, auto) nipkow@15392: txt {* finiteness of the two sets *} nipkow@15392: apply (rule_tac [2] B = "Pow (A)" in finite_subset) nipkow@15392: apply (rule_tac B = "Pow (insert x A)" in finite_subset) nipkow@15392: apply fast+ wenzelm@12396: done wenzelm@12396: nipkow@15392: text {* nipkow@15392: Main theorem: combinatorial statement about number of subsets of a set. nipkow@15392: *} wenzelm@12396: nipkow@15392: lemma n_sub_lemma: nipkow@15392: "!!A. finite A ==> card {B. B <= A & card B = k} = (card A choose k)" nipkow@15392: apply (induct k) nipkow@15392: apply (simp add: card_s_0_eq_empty, atomize) nipkow@15392: apply (rotate_tac -1, erule finite_induct) nipkow@15392: apply (simp_all (no_asm_simp) cong add: conj_cong nipkow@15392: add: card_s_0_eq_empty choose_deconstruct) nipkow@15392: apply (subst card_Un_disjoint) nipkow@15392: prefer 4 apply (force simp add: constr_bij) nipkow@15392: prefer 3 apply force nipkow@15392: prefer 2 apply (blast intro: finite_Pow_iff [THEN iffD2] nipkow@15392: finite_subset [of _ "Pow (insert x F)", standard]) nipkow@15392: apply (blast intro: finite_Pow_iff [THEN iffD2, THEN [2] finite_subset]) wenzelm@12396: done wenzelm@12396: nipkow@15392: theorem n_subsets: nipkow@15392: "finite A ==> card {B. B <= A & card B = k} = (card A choose k)" nipkow@15392: by (simp add: n_sub_lemma) nipkow@15392: nipkow@15392: nipkow@15392: subsection{* A fold functional for non-empty sets *} nipkow@15392: nipkow@15392: text{* Does not require start value. *} wenzelm@12396: nipkow@15392: consts nipkow@15392: foldSet1 :: "('a => 'a => 'a) => ('a set \ 'a) set" nipkow@15392: nipkow@15392: inductive "foldSet1 f" nipkow@15392: intros nipkow@15392: foldSet1_singletonI [intro]: "({a}, a) : foldSet1 f" nipkow@15392: foldSet1_insertI [intro]: nipkow@15392: "\ (A, x) : foldSet1 f; a \ A; A \ {} \ nipkow@15392: \ (insert a A, f a x) : foldSet1 f" wenzelm@12396: nipkow@15392: constdefs nipkow@15392: fold1 :: "('a => 'a => 'a) => 'a set => 'a" nipkow@15392: "fold1 f A == THE x. (A, x) : foldSet1 f" nipkow@15392: nipkow@15392: lemma foldSet1_nonempty: nipkow@15392: "(A, x) : foldSet1 f \ A \ {}" nipkow@15392: by(erule foldSet1.cases, simp_all) nipkow@15392: wenzelm@12396: nipkow@15392: inductive_cases empty_foldSet1E [elim!]: "({}, x) : foldSet1 f" nipkow@15392: nipkow@15392: lemma foldSet1_sing[iff]: "(({a},b) : foldSet1 f) = (a = b)" nipkow@15392: apply(rule iffI) nipkow@15392: prefer 2 apply fast nipkow@15392: apply (erule foldSet1.cases) nipkow@15392: apply blast nipkow@15392: apply (erule foldSet1.cases) nipkow@15392: apply blast nipkow@15392: apply blast nipkow@15376: done wenzelm@12396: nipkow@15392: lemma Diff1_foldSet1: nipkow@15392: "(A - {x}, y) : foldSet1 f ==> x: A ==> (A, f x y) : foldSet1 f" nipkow@15392: by (erule insert_Diff [THEN subst], rule foldSet1.intros, nipkow@15392: auto dest!:foldSet1_nonempty) wenzelm@12396: nipkow@15392: lemma foldSet1_imp_finite: "(A, x) : foldSet1 f ==> finite A" nipkow@15392: by (induct set: foldSet1) auto wenzelm@12396: nipkow@15392: lemma finite_nonempty_imp_foldSet1: nipkow@15392: "\ finite A; A \ {} \ \ EX x. (A, x) : foldSet1 f" nipkow@15392: by (induct set: Finites) auto nipkow@15376: nipkow@15392: lemma (in ACf) foldSet1_determ_aux: nipkow@15392: "!!A x y. \ card A < n; (A, x) : foldSet1 f; (A, y) : foldSet1 f \ \ y = x" nipkow@15392: proof (induct n) nipkow@15392: case 0 thus ?case by simp nipkow@15392: next nipkow@15392: case (Suc n) nipkow@15392: have IH: "!!A x y. \card A < n; (A, x) \ foldSet1 f; (A, y) \ foldSet1 f\ nipkow@15392: \ y = x" and card: "card A < Suc n" nipkow@15392: and Afoldx: "(A, x) \ foldSet1 f" and Afoldy: "(A, y) \ foldSet1 f" . nipkow@15392: from card have "card A < n \ card A = n" by arith nipkow@15392: thus ?case nipkow@15392: proof nipkow@15392: assume less: "card A < n" nipkow@15392: show ?thesis by(rule IH[OF less Afoldx Afoldy]) nipkow@15392: next nipkow@15392: assume cardA: "card A = n" nipkow@15392: show ?thesis nipkow@15392: proof (rule foldSet1.cases[OF Afoldx]) nipkow@15392: fix a assume "(A, x) = ({a}, a)" nipkow@15392: thus "y = x" using Afoldy by (simp add:foldSet1_sing) nipkow@15392: next nipkow@15392: fix Ax ax x' nipkow@15392: assume eq1: "(A, x) = (insert ax Ax, ax \ x')" nipkow@15392: and x': "(Ax, x') \ foldSet1 f" and notinx: "ax \ Ax" nipkow@15392: and Axnon: "Ax \ {}" nipkow@15392: hence A1: "A = insert ax Ax" and x: "x = ax \ x'" by auto nipkow@15392: show ?thesis nipkow@15392: proof (rule foldSet1.cases[OF Afoldy]) nipkow@15392: fix ay assume "(A, y) = ({ay}, ay)" nipkow@15392: thus ?thesis using eq1 x' Axnon notinx nipkow@15392: by (fastsimp simp:foldSet1_sing) nipkow@15392: next nipkow@15392: fix Ay ay y' nipkow@15392: assume eq2: "(A, y) = (insert ay Ay, ay \ y')" nipkow@15392: and y': "(Ay, y') \ foldSet1 f" and notiny: "ay \ Ay" nipkow@15392: and Aynon: "Ay \ {}" nipkow@15392: hence A2: "A = insert ay Ay" and y: "y = ay \ y'" by auto nipkow@15392: have finA: "finite A" by(rule foldSet1_imp_finite[OF Afoldx]) nipkow@15392: with cardA A1 notinx have less: "card Ax < n" by simp nipkow@15392: show ?thesis nipkow@15392: proof cases nipkow@15392: assume "ax = ay" nipkow@15392: then moreover have "Ax = Ay" using A1 A2 notinx notiny by auto nipkow@15392: ultimately show ?thesis using IH[OF less x'] y' eq1 eq2 by auto nipkow@15392: next nipkow@15392: assume diff: "ax \ ay" nipkow@15392: let ?B = "Ax - {ay}" nipkow@15392: have Ax: "Ax = insert ay ?B" and Ay: "Ay = insert ax ?B" nipkow@15392: using A1 A2 notinx notiny diff by(blast elim!:equalityE)+ nipkow@15392: show ?thesis nipkow@15392: proof cases nipkow@15392: assume "?B = {}" nipkow@15392: with Ax Ay show ?thesis using x' y' x y by(simp add:commute) nipkow@15392: next nipkow@15392: assume Bnon: "?B \ {}" nipkow@15392: moreover have "finite ?B" using finA A1 by simp nipkow@15392: ultimately obtain b where Bfoldb: "(?B,b) \ foldSet1 f" nipkow@15392: using finite_nonempty_imp_foldSet1 by(blast) nipkow@15392: moreover have ayinAx: "ay \ Ax" using Ax by(auto) nipkow@15392: ultimately have "(Ax,ay\b) \ foldSet1 f" by(rule Diff1_foldSet1) nipkow@15392: hence "ay\b = x'" by(rule IH[OF less x']) nipkow@15392: moreover have "ax\b = y'" nipkow@15392: proof (rule IH[OF _ y']) nipkow@15392: show "card Ay < n" using Ay cardA A1 notinx finA ayinAx nipkow@15392: by(auto simp:card_Diff1_less) nipkow@15392: next nipkow@15392: show "(Ay,ax\b) \ foldSet1 f" using Ay notinx Bfoldb Bnon nipkow@15392: by fastsimp nipkow@15392: qed nipkow@15392: ultimately show ?thesis using x y by(auto simp:AC) nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed nipkow@15392: qed wenzelm@12396: qed wenzelm@12396: nipkow@15392: nipkow@15392: lemma (in ACf) foldSet1_determ: nipkow@15392: "(A, x) : foldSet1 f ==> (A, y) : foldSet1 f ==> y = x" nipkow@15392: by (blast intro: foldSet1_determ_aux [rule_format]) nipkow@15392: nipkow@15392: lemma (in ACf) foldSet1_equality: "(A, y) : foldSet1 f ==> fold1 f A = y" nipkow@15392: by (unfold fold1_def) (blast intro: foldSet1_determ) nipkow@15392: nipkow@15392: lemma fold1_singleton: "fold1 f {a} = a" nipkow@15392: by (unfold fold1_def) blast wenzelm@12396: nipkow@15392: lemma (in ACf) foldSet1_insert_aux: "x \ A ==> A \ {} \ nipkow@15392: ((insert x A, v) : foldSet1 f) = nipkow@15392: (EX y. (A, y) : foldSet1 f & v = f x y)" nipkow@15392: apply auto nipkow@15392: apply (rule_tac A1 = A and f1 = f in finite_nonempty_imp_foldSet1 [THEN exE]) nipkow@15392: apply (fastsimp dest: foldSet1_imp_finite) nipkow@15392: apply blast nipkow@15392: apply (blast intro: foldSet1_determ) nipkow@15392: done nipkow@15376: nipkow@15392: lemma (in ACf) fold1_insert: nipkow@15392: "finite A ==> x \ A ==> A \ {} \ fold1 f (insert x A) = f x (fold1 f A)" nipkow@15392: apply (unfold fold1_def) nipkow@15392: apply (simp add: foldSet1_insert_aux) nipkow@15392: apply (rule the_equality) nipkow@15392: apply (auto intro: finite_nonempty_imp_foldSet1 nipkow@15392: cong add: conj_cong simp add: fold1_def [symmetric] foldSet1_equality) nipkow@15392: done nipkow@15376: nipkow@15392: locale ACIf = ACf + nipkow@15392: assumes idem: "x \ x = x" wenzelm@12396: nipkow@15392: lemma (in ACIf) fold1_insert2: nipkow@15392: assumes finA: "finite A" and nonA: "A \ {}" nipkow@15392: shows "fold1 f (insert a A) = f a (fold1 f A)" nipkow@15392: proof cases nipkow@15392: assume "a \ A" nipkow@15392: then obtain B where A: "A = insert a B" and disj: "a \ B" nipkow@15392: by(blast dest: mk_disjoint_insert) nipkow@15392: show ?thesis nipkow@15392: proof cases nipkow@15392: assume "B = {}" nipkow@15392: thus ?thesis using A by(simp add:idem fold1_singleton) nipkow@15392: next nipkow@15392: assume nonB: "B \ {}" nipkow@15392: from finA A have finB: "finite B" by(blast intro: finite_subset) nipkow@15392: have "fold1 f (insert a A) = fold1 f (insert a B)" using A by simp nipkow@15392: also have "\ = f a (fold1 f B)" nipkow@15392: using finB nonB disj by(simp add: fold1_insert) nipkow@15392: also have "\ = f a (fold1 f A)" nipkow@15392: using A finB nonB disj by(simp add:idem fold1_insert assoc[symmetric]) nipkow@15392: finally show ?thesis . nipkow@15392: qed nipkow@15392: next nipkow@15392: assume "a \ A" nipkow@15392: with finA nonA show ?thesis by(simp add:fold1_insert) nipkow@15392: qed nipkow@15392: nipkow@15376: nipkow@15392: text{* Now the recursion rules for definitions: *} nipkow@15392: nipkow@15392: lemma fold1_singleton_def: "g \ fold1 f \ g {a} = a" nipkow@15392: by(simp add:fold1_singleton) nipkow@15392: nipkow@15392: lemma (in ACf) fold1_insert_def: nipkow@15392: "\ g \ fold1 f; finite A; x \ A; A \ {} \ \ g(insert x A) = x \ (g A)" nipkow@15392: by(simp add:fold1_insert) nipkow@15392: nipkow@15392: lemma (in ACIf) fold1_insert2_def: nipkow@15392: "\ g \ fold1 f; finite A; A \ {} \ \ g(insert x A) = x \ (g A)" nipkow@15392: by(simp add:fold1_insert2) nipkow@15392: nipkow@15376: nipkow@15392: subsection{*Min and Max*} nipkow@15392: nipkow@15392: text{* As an application of @{text fold1} we define the minimal and nipkow@15392: maximal element of a (non-empty) set over a linear order. First we nipkow@15392: show that @{text min} and @{text max} are ACI: *} nipkow@15392: nipkow@15392: lemma ACf_min: "ACf(min :: 'a::linorder \ 'a \ 'a)" nipkow@15392: apply(rule ACf.intro) nipkow@15392: apply(auto simp:min_def) nipkow@15392: done nipkow@15392: nipkow@15392: lemma ACIf_min: "ACIf(min:: 'a::linorder \ 'a \ 'a)" nipkow@15392: apply(rule ACIf.intro[OF ACf_min]) nipkow@15392: apply(rule ACIf_axioms.intro) nipkow@15392: apply(auto simp:min_def) nipkow@15376: done nipkow@15376: nipkow@15392: lemma ACf_max: "ACf(max :: 'a::linorder \ 'a \ 'a)" nipkow@15392: apply(rule ACf.intro) nipkow@15392: apply(auto simp:max_def) nipkow@15392: done nipkow@15392: nipkow@15392: lemma ACIf_max: "ACIf(max:: 'a::linorder \ 'a \ 'a)" nipkow@15392: apply(rule ACIf.intro[OF ACf_max]) nipkow@15392: apply(rule ACIf_axioms.intro) nipkow@15392: apply(auto simp:max_def) nipkow@15376: done wenzelm@12396: nipkow@15392: text{* Now we make the definitions: *} nipkow@15392: nipkow@15392: constdefs nipkow@15392: Min :: "('a::linorder)set => 'a" nipkow@15392: "Min == fold1 min" nipkow@15392: nipkow@15392: Max :: "('a::linorder)set => 'a" nipkow@15392: "Max == fold1 max" nipkow@15392: nipkow@15402: text{* Now we instantiate the recursion equations and declare them nipkow@15392: simplification rules: *} nipkow@15392: nipkow@15392: declare nipkow@15392: fold1_singleton_def[OF Min_def, simp] nipkow@15392: ACIf.fold1_insert2_def[OF ACIf_min Min_def, simp] nipkow@15392: fold1_singleton_def[OF Max_def, simp] nipkow@15392: ACIf.fold1_insert2_def[OF ACIf_max Max_def, simp] nipkow@15392: nipkow@15392: text{* Now we prove some properties by induction: *} nipkow@15392: nipkow@15392: lemma Min_in [simp]: nipkow@15392: assumes a: "finite S" nipkow@15392: shows "S \ {} \ Min S \ S" nipkow@15392: using a nipkow@15392: proof induct nipkow@15392: case empty thus ?case by simp nipkow@15392: next nipkow@15392: case (insert x S) nipkow@15392: show ?case nipkow@15392: proof cases nipkow@15392: assume "S = {}" thus ?thesis by simp nipkow@15392: next nipkow@15392: assume "S \ {}" thus ?thesis using insert by (simp add:min_def) nipkow@15392: qed nipkow@15392: qed nipkow@15392: nipkow@15392: lemma Min_le [simp]: nipkow@15392: assumes a: "finite S" nipkow@15392: shows "\ S \ {}; x \ S \ \ Min S \ x" nipkow@15392: using a nipkow@15392: proof induct nipkow@15392: case empty thus ?case by simp nipkow@15392: next nipkow@15392: case (insert y S) nipkow@15392: show ?case nipkow@15392: proof cases nipkow@15392: assume "S = {}" thus ?thesis using insert by simp nipkow@15392: next nipkow@15392: assume "S \ {}" thus ?thesis using insert by (auto simp add:min_def) nipkow@15392: qed nipkow@15392: qed nipkow@15392: nipkow@15392: lemma Max_in [simp]: nipkow@15392: assumes a: "finite S" nipkow@15392: shows "S \ {} \ Max S \ S" nipkow@15392: using a nipkow@15392: proof induct nipkow@15392: case empty thus ?case by simp nipkow@15392: next nipkow@15392: case (insert x S) nipkow@15392: show ?case nipkow@15392: proof cases nipkow@15392: assume "S = {}" thus ?thesis by simp nipkow@15392: next nipkow@15392: assume "S \ {}" thus ?thesis using insert by (simp add:max_def) nipkow@15392: qed nipkow@15392: qed nipkow@15392: nipkow@15392: lemma Max_le [simp]: nipkow@15392: assumes a: "finite S" nipkow@15392: shows "\ S \ {}; x \ S \ \ x \ Max S" nipkow@15392: using a nipkow@15392: proof induct nipkow@15392: case empty thus ?case by simp nipkow@15392: next nipkow@15392: case (insert y S) nipkow@15392: show ?case nipkow@15392: proof cases nipkow@15392: assume "S = {}" thus ?thesis using insert by simp nipkow@15392: next nipkow@15392: assume "S \ {}" thus ?thesis using insert by (auto simp add:max_def) nipkow@15392: qed nipkow@15392: qed nipkow@15392: wenzelm@12396: nipkow@15042: end