diff --git a/src/HOL/Finite_Set.thy b/src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy +++ b/src/HOL/Finite_Set.thy @@ -1,2266 +1,2288 @@ (* Title: HOL/Finite_Set.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad Author: Andrei Popescu *) section \Finite sets\ theory Finite_Set imports Product_Type Sum_Type Fields begin subsection \Predicate for finite sets\ context notes [[inductive_internals]] begin inductive finite :: "'a set \ bool" where emptyI [simp, intro!]: "finite {}" | insertI [simp, intro!]: "finite A \ finite (insert a A)" end simproc_setup finite_Collect ("finite (Collect P)") = \K Set_Comprehension_Pointfree.simproc\ declare [[simproc del: finite_Collect]] lemma finite_induct [case_names empty insert, induct set: finite]: \ \Discharging \x \ F\ entails extra work.\ assumes "finite F" assumes "P {}" and insert: "\x F. finite F \ x \ F \ P F \ P (insert x F)" shows "P F" using \finite F\ proof induct show "P {}" by fact next fix x F assume F: "finite F" and P: "P F" show "P (insert x F)" proof cases assume "x \ F" then have "insert x F = F" by (rule insert_absorb) with P show ?thesis by (simp only:) next assume "x \ F" from F this P show ?thesis by (rule insert) qed qed lemma infinite_finite_induct [case_names infinite empty insert]: assumes infinite: "\A. \ finite A \ P A" and empty: "P {}" and insert: "\x F. finite F \ x \ F \ P F \ P (insert x F)" shows "P A" proof (cases "finite A") case False with infinite show ?thesis . next case True then show ?thesis by (induct A) (fact empty insert)+ qed subsubsection \Choice principles\ lemma ex_new_if_finite: \ \does not depend on def of finite at all\ assumes "\ finite (UNIV :: 'a set)" and "finite A" shows "\a::'a. a \ A" proof - from assms have "A \ UNIV" by blast then show ?thesis by blast qed text \A finite choice principle. Does not need the SOME choice operator.\ lemma finite_set_choice: "finite A \ \x\A. \y. P x y \ \f. \x\A. P x (f x)" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert a A) then obtain f b where f: "\x\A. P x (f x)" and ab: "P a b" by auto show ?case (is "\f. ?P f") proof show "?P (\x. if x = a then b else f x)" using f ab by auto qed qed subsubsection \Finite sets are the images of initial segments of natural numbers\ lemma finite_imp_nat_seg_image_inj_on: assumes "finite A" shows "\(n::nat) f. A = f ` {i. i < n} \ inj_on f {i. i < n}" using assms proof induct case empty show ?case proof show "\f. {} = f ` {i::nat. i < 0} \ inj_on f {i. i < 0}" by simp qed next case (insert a A) have notinA: "a \ A" by fact from insert.hyps obtain n f where "A = f ` {i::nat. i < n}" "inj_on f {i. i < n}" by blast then have "insert a A = f(n:=a) ` {i. i < Suc n}" and "inj_on (f(n:=a)) {i. i < Suc n}" using notinA by (auto simp add: image_def Ball_def inj_on_def less_Suc_eq) then show ?case by blast qed lemma nat_seg_image_imp_finite: "A = f ` {i::nat. i < n} \ finite A" proof (induct n arbitrary: A) case 0 then show ?case by simp next case (Suc n) let ?B = "f ` {i. i < n}" have finB: "finite ?B" by (rule Suc.hyps[OF refl]) show ?case proof (cases "\k (\n f. A = f ` {i::nat. i < n})" by (blast intro: nat_seg_image_imp_finite dest: finite_imp_nat_seg_image_inj_on) lemma finite_imp_inj_to_nat_seg: assumes "finite A" shows "\f n. f ` A = {i::nat. i < n} \ inj_on f A" proof - from finite_imp_nat_seg_image_inj_on [OF \finite A\] obtain f and n :: nat where bij: "bij_betw f {i. i ?f ` A = {i. i k}" by (simp add: le_eq_less_or_eq Collect_disj_eq) subsection \Finiteness and common set operations\ lemma rev_finite_subset: "finite B \ A \ B \ finite A" proof (induct arbitrary: A rule: finite_induct) case empty then show ?case by simp next case (insert x F A) have A: "A \ insert x F" and r: "A - {x} \ F \ finite (A - {x})" by fact+ show "finite A" proof cases assume x: "x \ A" with A have "A - {x} \ F" by (simp add: subset_insert_iff) with r have "finite (A - {x})" . then have "finite (insert x (A - {x}))" .. also have "insert x (A - {x}) = A" using x by (rule insert_Diff) finally show ?thesis . next show ?thesis when "A \ F" using that by fact assume "x \ A" with A show "A \ F" by (simp add: subset_insert_iff) qed qed lemma finite_subset: "A \ B \ finite B \ finite A" by (rule rev_finite_subset) lemma finite_UnI: assumes "finite F" and "finite G" shows "finite (F \ G)" using assms by induct simp_all lemma finite_Un [iff]: "finite (F \ G) \ finite F \ finite G" by (blast intro: finite_UnI finite_subset [of _ "F \ G"]) lemma finite_insert [simp]: "finite (insert a A) \ finite A" proof - have "finite {a} \ finite A \ finite A" by simp then have "finite ({a} \ A) \ finite A" by (simp only: finite_Un) then show ?thesis by simp qed lemma finite_Int [simp, intro]: "finite F \ finite G \ finite (F \ G)" by (blast intro: finite_subset) lemma finite_Collect_conjI [simp, intro]: "finite {x. P x} \ finite {x. Q x} \ finite {x. P x \ Q x}" by (simp add: Collect_conj_eq) lemma finite_Collect_disjI [simp]: "finite {x. P x \ Q x} \ finite {x. P x} \ finite {x. Q x}" by (simp add: Collect_disj_eq) lemma finite_Diff [simp, intro]: "finite A \ finite (A - B)" by (rule finite_subset, rule Diff_subset) lemma finite_Diff2 [simp]: assumes "finite B" shows "finite (A - B) \ finite A" proof - have "finite A \ finite ((A - B) \ (A \ B))" by (simp add: Un_Diff_Int) also have "\ \ finite (A - B)" using \finite B\ by simp finally show ?thesis .. qed lemma finite_Diff_insert [iff]: "finite (A - insert a B) \ finite (A - B)" proof - have "finite (A - B) \ finite (A - B - {a})" by simp moreover have "A - insert a B = A - B - {a}" by auto ultimately show ?thesis by simp qed lemma finite_compl [simp]: "finite (A :: 'a set) \ finite (- A) \ finite (UNIV :: 'a set)" by (simp add: Compl_eq_Diff_UNIV) lemma finite_Collect_not [simp]: "finite {x :: 'a. P x} \ finite {x. \ P x} \ finite (UNIV :: 'a set)" by (simp add: Collect_neg_eq) lemma finite_Union [simp, intro]: "finite A \ (\M. M \ A \ finite M) \ finite (\A)" by (induct rule: finite_induct) simp_all lemma finite_UN_I [intro]: "finite A \ (\a. a \ A \ finite (B a)) \ finite (\a\A. B a)" by (induct rule: finite_induct) simp_all lemma finite_UN [simp]: "finite A \ finite (\(B ` A)) \ (\x\A. finite (B x))" by (blast intro: finite_subset) lemma finite_Inter [intro]: "\A\M. finite A \ finite (\M)" by (blast intro: Inter_lower finite_subset) lemma finite_INT [intro]: "\x\I. finite (A x) \ finite (\x\I. A x)" by (blast intro: INT_lower finite_subset) lemma finite_imageI [simp, intro]: "finite F \ finite (h ` F)" by (induct rule: finite_induct) simp_all lemma finite_image_set [simp]: "finite {x. P x} \ finite {f x |x. P x}" by (simp add: image_Collect [symmetric]) lemma finite_image_set2: "finite {x. P x} \ finite {y. Q y} \ finite {f x y |x y. P x \ Q y}" by (rule finite_subset [where B = "\x \ {x. P x}. \y \ {y. Q y}. {f x y}"]) auto lemma finite_imageD: assumes "finite (f ` A)" and "inj_on f A" shows "finite A" using assms proof (induct "f ` A" arbitrary: A) case empty then show ?case by simp next case (insert x B) then have B_A: "insert x B = f ` A" by simp then obtain y where "x = f y" and "y \ A" by blast from B_A \x \ B\ have "B = f ` A - {x}" by blast with B_A \x \ B\ \x = f y\ \inj_on f A\ \y \ A\ have "B = f ` (A - {y})" by (simp add: inj_on_image_set_diff) moreover from \inj_on f A\ have "inj_on f (A - {y})" by (rule inj_on_diff) ultimately have "finite (A - {y})" by (rule insert.hyps) then show "finite A" by simp qed lemma finite_image_iff: "inj_on f A \ finite (f ` A) \ finite A" using finite_imageD by blast lemma finite_surj: "finite A \ B \ f ` A \ finite B" by (erule finite_subset) (rule finite_imageI) lemma finite_range_imageI: "finite (range g) \ finite (range (\x. f (g x)))" by (drule finite_imageI) (simp add: range_composition) lemma finite_subset_image: assumes "finite B" shows "B \ f ` A \ \C\A. finite C \ B = f ` C" using assms proof induct case empty then show ?case by simp next case insert then show ?case by (clarsimp simp del: image_insert simp add: image_insert [symmetric]) blast qed lemma all_subset_image: "(\B. B \ f ` A \ P B) \ (\B. B \ A \ P(f ` B))" by (safe elim!: subset_imageE) (use image_mono in \blast+\) (* slow *) lemma all_finite_subset_image: "(\B. finite B \ B \ f ` A \ P B) \ (\B. finite B \ B \ A \ P (f ` B))" proof safe fix B :: "'a set" assume B: "finite B" "B \ f ` A" and P: "\B. finite B \ B \ A \ P (f ` B)" show "P B" using finite_subset_image [OF B] P by blast qed blast lemma ex_finite_subset_image: "(\B. finite B \ B \ f ` A \ P B) \ (\B. finite B \ B \ A \ P (f ` B))" proof safe fix B :: "'a set" assume B: "finite B" "B \ f ` A" and "P B" show "\B. finite B \ B \ A \ P (f ` B)" using finite_subset_image [OF B] \P B\ by blast qed blast lemma finite_vimage_IntI: "finite F \ inj_on h A \ finite (h -` F \ A)" proof (induct rule: finite_induct) case (insert x F) then show ?case by (simp add: vimage_insert [of h x F] finite_subset [OF inj_on_vimage_singleton] Int_Un_distrib2) qed simp lemma finite_finite_vimage_IntI: assumes "finite F" and "\y. y \ F \ finite ((h -` {y}) \ A)" shows "finite (h -` F \ A)" proof - have *: "h -` F \ A = (\ y\F. (h -` {y}) \ A)" by blast show ?thesis by (simp only: * assms finite_UN_I) qed lemma finite_vimageI: "finite F \ inj h \ finite (h -` F)" using finite_vimage_IntI[of F h UNIV] by auto lemma finite_vimageD': "finite (f -` A) \ A \ range f \ finite A" by (auto simp add: subset_image_iff intro: finite_subset[rotated]) lemma finite_vimageD: "finite (h -` F) \ surj h \ finite F" by (auto dest: finite_vimageD') lemma finite_vimage_iff: "bij h \ finite (h -` F) \ finite F" unfolding bij_def by (auto elim: finite_vimageD finite_vimageI) lemma finite_Collect_bex [simp]: assumes "finite A" shows "finite {x. \y\A. Q x y} \ (\y\A. finite {x. Q x y})" proof - have "{x. \y\A. Q x y} = (\y\A. {x. Q x y})" by auto with assms show ?thesis by simp qed lemma finite_Collect_bounded_ex [simp]: assumes "finite {y. P y}" shows "finite {x. \y. P y \ Q x y} \ (\y. P y \ finite {x. Q x y})" proof - have "{x. \y. P y \ Q x y} = (\y\{y. P y}. {x. Q x y})" by auto with assms show ?thesis by simp qed lemma finite_Plus: "finite A \ finite B \ finite (A <+> B)" by (simp add: Plus_def) lemma finite_PlusD: fixes A :: "'a set" and B :: "'b set" assumes fin: "finite (A <+> B)" shows "finite A" "finite B" proof - have "Inl ` A \ A <+> B" by auto then have "finite (Inl ` A :: ('a + 'b) set)" using fin by (rule finite_subset) then show "finite A" by (rule finite_imageD) (auto intro: inj_onI) next have "Inr ` B \ A <+> B" by auto then have "finite (Inr ` B :: ('a + 'b) set)" using fin by (rule finite_subset) then show "finite B" by (rule finite_imageD) (auto intro: inj_onI) qed lemma finite_Plus_iff [simp]: "finite (A <+> B) \ finite A \ finite B" by (auto intro: finite_PlusD finite_Plus) lemma finite_Plus_UNIV_iff [simp]: "finite (UNIV :: ('a + 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set)" by (subst UNIV_Plus_UNIV [symmetric]) (rule finite_Plus_iff) lemma finite_SigmaI [simp, intro]: "finite A \ (\a. a\A \ finite (B a)) \ finite (SIGMA a:A. B a)" unfolding Sigma_def by blast lemma finite_SigmaI2: assumes "finite {x\A. B x \ {}}" and "\a. a \ A \ finite (B a)" shows "finite (Sigma A B)" proof - from assms have "finite (Sigma {x\A. B x \ {}} B)" by auto also have "Sigma {x:A. B x \ {}} B = Sigma A B" by auto finally show ?thesis . qed lemma finite_cartesian_product: "finite A \ finite B \ finite (A \ B)" by (rule finite_SigmaI) lemma finite_Prod_UNIV: "finite (UNIV :: 'a set) \ finite (UNIV :: 'b set) \ finite (UNIV :: ('a \ 'b) set)" by (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product) lemma finite_cartesian_productD1: assumes "finite (A \ B)" and "B \ {}" shows "finite A" proof - from assms obtain n f where "A \ B = f ` {i::nat. i < n}" by (auto simp add: finite_conv_nat_seg_image) then have "fst ` (A \ B) = fst ` f ` {i::nat. i < n}" by simp with \B \ {}\ have "A = (fst \ f) ` {i::nat. i < n}" by (simp add: image_comp) then have "\n f. A = f ` {i::nat. i < n}" by blast then show ?thesis by (auto simp add: finite_conv_nat_seg_image) qed lemma finite_cartesian_productD2: assumes "finite (A \ B)" and "A \ {}" shows "finite B" proof - from assms obtain n f where "A \ B = f ` {i::nat. i < n}" by (auto simp add: finite_conv_nat_seg_image) then have "snd ` (A \ B) = snd ` f ` {i::nat. i < n}" by simp with \A \ {}\ have "B = (snd \ f) ` {i::nat. i < n}" by (simp add: image_comp) then have "\n f. B = f ` {i::nat. i < n}" by blast then show ?thesis by (auto simp add: finite_conv_nat_seg_image) qed lemma finite_cartesian_product_iff: "finite (A \ B) \ (A = {} \ B = {} \ (finite A \ finite B))" by (auto dest: finite_cartesian_productD1 finite_cartesian_productD2 finite_cartesian_product) lemma finite_prod: "finite (UNIV :: ('a \ 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set)" using finite_cartesian_product_iff[of UNIV UNIV] by simp lemma finite_Pow_iff [iff]: "finite (Pow A) \ finite A" proof assume "finite (Pow A)" then have "finite ((\x. {x}) ` A)" by (blast intro: finite_subset) (* somewhat slow *) then show "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp next assume "finite A" then show "finite (Pow A)" by induct (simp_all add: Pow_insert) qed corollary finite_Collect_subsets [simp, intro]: "finite A \ finite {B. B \ A}" by (simp add: Pow_def [symmetric]) lemma finite_set: "finite (UNIV :: 'a set set) \ finite (UNIV :: 'a set)" by (simp only: finite_Pow_iff Pow_UNIV[symmetric]) lemma finite_UnionD: "finite (\A) \ finite A" by (blast intro: finite_subset [OF subset_Pow_Union]) lemma finite_bind: assumes "finite S" assumes "\x \ S. finite (f x)" shows "finite (Set.bind S f)" using assms by (simp add: bind_UNION) lemma finite_filter [simp]: "finite S \ finite (Set.filter P S)" unfolding Set.filter_def by simp lemma finite_set_of_finite_funs: assumes "finite A" "finite B" shows "finite {f. \x. (x \ A \ f x \ B) \ (x \ A \ f x = d)}" (is "finite ?S") proof - let ?F = "\f. {(a,b). a \ A \ b = f a}" have "?F ` ?S \ Pow(A \ B)" by auto from finite_subset[OF this] assms have 1: "finite (?F ` ?S)" by simp have 2: "inj_on ?F ?S" by (fastforce simp add: inj_on_def set_eq_iff fun_eq_iff) (* somewhat slow *) show ?thesis by (rule finite_imageD [OF 1 2]) qed lemma not_finite_existsD: assumes "\ finite {a. P a}" shows "\a. P a" proof (rule classical) assume "\ ?thesis" with assms show ?thesis by auto qed subsection \Further induction rules on finite sets\ lemma finite_ne_induct [case_names singleton insert, consumes 2]: assumes "finite F" and "F \ {}" assumes "\x. P {x}" and "\x F. finite F \ F \ {} \ x \ F \ P F \ P (insert x F)" shows "P F" using assms proof induct case empty then show ?case by simp next case (insert x F) then show ?case by cases auto qed lemma finite_subset_induct [consumes 2, case_names empty insert]: assumes "finite F" and "F \ A" and empty: "P {}" and insert: "\a F. finite F \ a \ A \ a \ F \ P F \ P (insert a F)" shows "P F" using \finite F\ \F \ A\ proof induct show "P {}" by fact next fix x F assume "finite F" and "x \ F" and P: "F \ A \ P F" and i: "insert x F \ A" show "P (insert x F)" proof (rule insert) from i show "x \ A" by blast from i have "F \ A" by blast with P show "P F" . show "finite F" by fact show "x \ F" by fact qed qed lemma finite_empty_induct: assumes "finite A" and "P A" and remove: "\a A. finite A \ a \ A \ P A \ P (A - {a})" shows "P {}" proof - have "P (A - B)" if "B \ A" for B :: "'a set" proof - from \finite A\ that have "finite B" by (rule rev_finite_subset) from this \B \ A\ show "P (A - B)" proof induct case empty from \P A\ show ?case by simp next case (insert b B) have "P (A - B - {b})" proof (rule remove) from \finite A\ show "finite (A - B)" by induct auto from insert show "b \ A - B" by simp from insert show "P (A - B)" by simp qed also have "A - B - {b} = A - insert b B" by (rule Diff_insert [symmetric]) finally show ?case . qed qed then have "P (A - A)" by blast then show ?thesis by simp qed lemma finite_update_induct [consumes 1, case_names const update]: assumes finite: "finite {a. f a \ c}" and const: "P (\a. c)" and update: "\a b f. finite {a. f a \ c} \ f a = c \ b \ c \ P f \ P (f(a := b))" shows "P f" using finite proof (induct "{a. f a \ c}" arbitrary: f) case empty with const show ?case by simp next case (insert a A) then have "A = {a'. (f(a := c)) a' \ c}" and "f a \ c" by auto with \finite A\ have "finite {a'. (f(a := c)) a' \ c}" by simp have "(f(a := c)) a = c" by simp from insert \A = {a'. (f(a := c)) a' \ c}\ have "P (f(a := c))" by simp with \finite {a'. (f(a := c)) a' \ c}\ \(f(a := c)) a = c\ \f a \ c\ have "P ((f(a := c))(a := f a))" by (rule update) then show ?case by simp qed lemma finite_subset_induct' [consumes 2, case_names empty insert]: assumes "finite F" and "F \ A" and empty: "P {}" and insert: "\a F. \finite F; a \ A; F \ A; a \ F; P F \ \ P (insert a F)" shows "P F" using assms(1,2) proof induct show "P {}" by fact next fix x F assume "finite F" and "x \ F" and P: "F \ A \ P F" and i: "insert x F \ A" show "P (insert x F)" proof (rule insert) from i show "x \ A" by blast from i have "F \ A" by blast with P show "P F" . show "finite F" by fact show "x \ F" by fact show "F \ A" by fact qed qed subsection \Class \finite\\ class finite = assumes finite_UNIV: "finite (UNIV :: 'a set)" begin lemma finite [simp]: "finite (A :: 'a set)" by (rule subset_UNIV finite_UNIV finite_subset)+ lemma finite_code [code]: "finite (A :: 'a set) \ True" by simp end instance prod :: (finite, finite) finite by standard (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite) lemma inj_graph: "inj (\f. {(x, y). y = f x})" by (rule inj_onI) (auto simp add: set_eq_iff fun_eq_iff) instance "fun" :: (finite, finite) finite proof show "finite (UNIV :: ('a \ 'b) set)" proof (rule finite_imageD) let ?graph = "\f::'a \ 'b. {(x, y). y = f x}" have "range ?graph \ Pow UNIV" by simp moreover have "finite (Pow (UNIV :: ('a * 'b) set))" by (simp only: finite_Pow_iff finite) ultimately show "finite (range ?graph)" by (rule finite_subset) show "inj ?graph" by (rule inj_graph) qed qed instance bool :: finite by standard (simp add: UNIV_bool) instance set :: (finite) finite by standard (simp only: Pow_UNIV [symmetric] finite_Pow_iff finite) instance unit :: finite by standard (simp add: UNIV_unit) instance sum :: (finite, finite) finite by standard (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite) subsection \A basic fold functional for finite sets\ text \The intended behaviour is \fold f z {x\<^sub>1, \, x\<^sub>n} = f x\<^sub>1 (\ (f x\<^sub>n z)\)\ if \f\ is ``left-commutative'': \ locale comp_fun_commute = fixes f :: "'a \ 'b \ 'b" assumes comp_fun_commute: "f y \ f x = f x \ f y" begin lemma fun_left_comm: "f y (f x z) = f x (f y z)" using comp_fun_commute by (simp add: fun_eq_iff) lemma commute_left_comp: "f y \ (f x \ g) = f x \ (f y \ g)" by (simp add: o_assoc comp_fun_commute) end inductive fold_graph :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b \ bool" for f :: "'a \ 'b \ 'b" and z :: 'b where emptyI [intro]: "fold_graph f z {} z" | insertI [intro]: "x \ A \ fold_graph f z A y \ fold_graph f z (insert x A) (f x y)" inductive_cases empty_fold_graphE [elim!]: "fold_graph f z {} x" lemma fold_graph_closed_lemma: "fold_graph f z A x \ x \ B" if "fold_graph g z A x" "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" using that(1-3) proof (induction rule: fold_graph.induct) case (insertI x A y) have "fold_graph f z A y" "y \ B" unfolding atomize_conj by (rule insertI.IH) (auto intro: insertI.prems) then have "g x y \ B" and f_eq: "f x y = g x y" by (auto simp: insertI.prems) moreover have "fold_graph f z (insert x A) (f x y)" by (rule fold_graph.insertI; fact) ultimately show ?case by (simp add: f_eq) qed (auto intro!: that) lemma fold_graph_closed_eq: "fold_graph f z A = fold_graph g z A" if "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" using fold_graph_closed_lemma[of f z A _ B g] fold_graph_closed_lemma[of g z A _ B f] that by auto definition fold :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" where "fold f z A = (if finite A then (THE y. fold_graph f z A y) else z)" lemma fold_closed_eq: "fold f z A = fold g z A" if "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" unfolding Finite_Set.fold_def by (subst fold_graph_closed_eq[where B=B and g=g]) (auto simp: that) text \ A tempting alternative for the definiens is \<^term>\if finite A then THE y. fold_graph f z A y else e\. It allows the removal of finiteness assumptions from the theorems \fold_comm\, \fold_reindex\ and \fold_distrib\. The proofs become ugly. It is not worth the effort. (???) \ lemma finite_imp_fold_graph: "finite A \ \x. fold_graph f z A x" by (induct rule: finite_induct) auto subsubsection \From \<^const>\fold_graph\ to \<^term>\fold\\ context comp_fun_commute begin lemma fold_graph_finite: assumes "fold_graph f z A y" shows "finite A" using assms by induct simp_all lemma fold_graph_insertE_aux: "fold_graph f z A y \ a \ A \ \y'. y = f a y' \ fold_graph f z (A - {a}) y'" proof (induct set: fold_graph) case emptyI then show ?case by simp next case (insertI x A y) show ?case proof (cases "x = a") case True with insertI show ?thesis by auto next case False then obtain y' where y: "y = f a y'" and y': "fold_graph f z (A - {a}) y'" using insertI by auto have "f x y = f a (f x y')" unfolding y by (rule fun_left_comm) moreover have "fold_graph f z (insert x A - {a}) (f x y')" using y' and \x \ a\ and \x \ A\ by (simp add: insert_Diff_if fold_graph.insertI) ultimately show ?thesis by fast qed qed lemma fold_graph_insertE: assumes "fold_graph f z (insert x A) v" and "x \ A" obtains y where "v = f x y" and "fold_graph f z A y" using assms by (auto dest: fold_graph_insertE_aux [OF _ insertI1]) lemma fold_graph_determ: "fold_graph f z A x \ fold_graph f z A y \ y = x" proof (induct arbitrary: y set: fold_graph) case emptyI then show ?case by fast next case (insertI x A y v) from \fold_graph f z (insert x A) v\ and \x \ A\ obtain y' where "v = f x y'" and "fold_graph f z A y'" by (rule fold_graph_insertE) from \fold_graph f z A y'\ have "y' = y" by (rule insertI) with \v = f x y'\ show "v = f x y" by simp qed lemma fold_equality: "fold_graph f z A y \ fold f z A = y" by (cases "finite A") (auto simp add: fold_def intro: fold_graph_determ dest: fold_graph_finite) lemma fold_graph_fold: assumes "finite A" shows "fold_graph f z A (fold f z A)" proof - from assms have "\x. fold_graph f z A x" by (rule finite_imp_fold_graph) moreover note fold_graph_determ ultimately have "\!x. fold_graph f z A x" by (rule ex_ex1I) then have "fold_graph f z A (The (fold_graph f z A))" by (rule theI') with assms show ?thesis by (simp add: fold_def) qed text \The base case for \fold\:\ lemma (in -) fold_infinite [simp]: "\ finite A \ fold f z A = z" by (auto simp: fold_def) lemma (in -) fold_empty [simp]: "fold f z {} = z" by (auto simp: fold_def) text \The various recursion equations for \<^const>\fold\:\ lemma fold_insert [simp]: assumes "finite A" and "x \ A" shows "fold f z (insert x A) = f x (fold f z A)" proof (rule fold_equality) fix z from \finite A\ have "fold_graph f z A (fold f z A)" by (rule fold_graph_fold) with \x \ A\ have "fold_graph f z (insert x A) (f x (fold f z A))" by (rule fold_graph.insertI) then show "fold_graph f z (insert x A) (f x (fold f z A))" by simp qed declare (in -) empty_fold_graphE [rule del] fold_graph.intros [rule del] \ \No more proofs involve these.\ lemma fold_fun_left_comm: "finite A \ f x (fold f z A) = fold f (f x z) A" proof (induct rule: finite_induct) case empty then show ?case by simp next case insert then show ?case by (simp add: fun_left_comm [of x]) qed lemma fold_insert2: "finite A \ x \ A \ fold f z (insert x A) = fold f (f x z) A" by (simp add: fold_fun_left_comm) lemma fold_rec: assumes "finite A" and "x \ A" shows "fold f z A = f x (fold f z (A - {x}))" proof - have A: "A = insert x (A - {x})" using \x \ A\ by blast then have "fold f z A = fold f z (insert x (A - {x}))" by simp also have "\ = f x (fold f z (A - {x}))" by (rule fold_insert) (simp add: \finite A\)+ finally show ?thesis . qed lemma fold_insert_remove: assumes "finite A" shows "fold f z (insert x A) = f x (fold f z (A - {x}))" proof - from \finite A\ have "finite (insert x A)" by auto moreover have "x \ insert x A" by auto ultimately have "fold f z (insert x A) = f x (fold f z (insert x A - {x}))" by (rule fold_rec) then show ?thesis by simp qed lemma fold_set_union_disj: assumes "finite A" "finite B" "A \ B = {}" shows "Finite_Set.fold f z (A \ B) = Finite_Set.fold f (Finite_Set.fold f z A) B" using assms(2,1,3) by induct simp_all end text \Other properties of \<^const>\fold\:\ lemma fold_image: assumes "inj_on g A" shows "fold f z (g ` A) = fold (f \ g) z A" proof (cases "finite A") case False with assms show ?thesis by (auto dest: finite_imageD simp add: fold_def) next case True have "fold_graph f z (g ` A) = fold_graph (f \ g) z A" proof fix w show "fold_graph f z (g ` A) w \ fold_graph (f \ g) z A w" (is "?P \ ?Q") proof assume ?P then show ?Q using assms proof (induct "g ` A" w arbitrary: A) case emptyI then show ?case by (auto intro: fold_graph.emptyI) next case (insertI x A r B) from \inj_on g B\ \x \ A\ \insert x A = image g B\ obtain x' A' where "x' \ A'" and [simp]: "B = insert x' A'" "x = g x'" "A = g ` A'" by (rule inj_img_insertE) from insertI.prems have "fold_graph (f \ g) z A' r" by (auto intro: insertI.hyps) with \x' \ A'\ have "fold_graph (f \ g) z (insert x' A') ((f \ g) x' r)" by (rule fold_graph.insertI) then show ?case by simp qed next assume ?Q then show ?P using assms proof induct case emptyI then show ?case by (auto intro: fold_graph.emptyI) next case (insertI x A r) from \x \ A\ insertI.prems have "g x \ g ` A" by auto moreover from insertI have "fold_graph f z (g ` A) r" by simp ultimately have "fold_graph f z (insert (g x) (g ` A)) (f (g x) r)" by (rule fold_graph.insertI) then show ?case by simp qed qed qed with True assms show ?thesis by (auto simp add: fold_def) qed lemma fold_cong: assumes "comp_fun_commute f" "comp_fun_commute g" and "finite A" and cong: "\x. x \ A \ f x = g x" and "s = t" and "A = B" shows "fold f s A = fold g t B" proof - have "fold f s A = fold g s A" using \finite A\ cong proof (induct A) case empty then show ?case by simp next case insert interpret f: comp_fun_commute f by (fact \comp_fun_commute f\) interpret g: comp_fun_commute g by (fact \comp_fun_commute g\) from insert show ?case by simp qed with assms show ?thesis by simp qed text \A simplified version for idempotent functions:\ locale comp_fun_idem = comp_fun_commute + assumes comp_fun_idem: "f x \ f x = f x" begin lemma fun_left_idem: "f x (f x z) = f x z" using comp_fun_idem by (simp add: fun_eq_iff) lemma fold_insert_idem: assumes fin: "finite A" shows "fold f z (insert x A) = f x (fold f z A)" proof cases assume "x \ A" then obtain B where "A = insert x B" and "x \ B" by (rule set_insert) then show ?thesis using assms by (simp add: comp_fun_idem fun_left_idem) next assume "x \ A" then show ?thesis using assms by simp qed declare fold_insert [simp del] fold_insert_idem [simp] lemma fold_insert_idem2: "finite A \ fold f z (insert x A) = fold f (f x z) A" by (simp add: fold_fun_left_comm) end subsubsection \Liftings to \comp_fun_commute\ etc.\ lemma (in comp_fun_commute) comp_comp_fun_commute: "comp_fun_commute (f \ g)" by standard (simp_all add: comp_fun_commute) lemma (in comp_fun_idem) comp_comp_fun_idem: "comp_fun_idem (f \ g)" by (rule comp_fun_idem.intro, rule comp_comp_fun_commute, unfold_locales) (simp_all add: comp_fun_idem) lemma (in comp_fun_commute) comp_fun_commute_funpow: "comp_fun_commute (\x. f x ^^ g x)" proof show "f y ^^ g y \ f x ^^ g x = f x ^^ g x \ f y ^^ g y" for x y proof (cases "x = y") case True then show ?thesis by simp next case False show ?thesis proof (induct "g x" arbitrary: g) case 0 then show ?case by simp next case (Suc n g) have hyp1: "f y ^^ g y \ f x = f x \ f y ^^ g y" proof (induct "g y" arbitrary: g) case 0 then show ?case by simp next case (Suc n g) define h where "h z = g z - 1" for z with Suc have "n = h y" by simp with Suc have hyp: "f y ^^ h y \ f x = f x \ f y ^^ h y" by auto from Suc h_def have "g y = Suc (h y)" by simp then show ?case by (simp add: comp_assoc hyp) (simp add: o_assoc comp_fun_commute) qed define h where "h z = (if z = x then g x - 1 else g z)" for z with Suc have "n = h x" by simp with Suc have "f y ^^ h y \ f x ^^ h x = f x ^^ h x \ f y ^^ h y" by auto with False h_def have hyp2: "f y ^^ g y \ f x ^^ h x = f x ^^ h x \ f y ^^ g y" by simp from Suc h_def have "g x = Suc (h x)" by simp then show ?case by (simp del: funpow.simps add: funpow_Suc_right o_assoc hyp2) (simp add: comp_assoc hyp1) qed qed qed subsubsection \Expressing set operations via \<^const>\fold\\ lemma comp_fun_commute_const: "comp_fun_commute (\_. f)" by standard rule lemma comp_fun_idem_insert: "comp_fun_idem insert" by standard auto lemma comp_fun_idem_remove: "comp_fun_idem Set.remove" by standard auto lemma (in semilattice_inf) comp_fun_idem_inf: "comp_fun_idem inf" by standard (auto simp add: inf_left_commute) lemma (in semilattice_sup) comp_fun_idem_sup: "comp_fun_idem sup" by standard (auto simp add: sup_left_commute) lemma union_fold_insert: assumes "finite A" shows "A \ B = fold insert B A" proof - interpret comp_fun_idem insert by (fact comp_fun_idem_insert) from \finite A\ show ?thesis by (induct A arbitrary: B) simp_all qed lemma minus_fold_remove: assumes "finite A" shows "B - A = fold Set.remove B A" proof - interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove) from \finite A\ have "fold Set.remove B A = B - A" by (induct A arbitrary: B) auto (* slow *) then show ?thesis .. qed lemma comp_fun_commute_filter_fold: "comp_fun_commute (\x A'. if P x then Set.insert x A' else A')" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show ?thesis by standard (auto simp: fun_eq_iff) qed lemma Set_filter_fold: assumes "finite A" shows "Set.filter P A = fold (\x A'. if P x then Set.insert x A' else A') {} A" using assms by induct (auto simp add: Set.filter_def comp_fun_commute.fold_insert[OF comp_fun_commute_filter_fold]) lemma inter_Set_filter: assumes "finite B" shows "A \ B = Set.filter (\x. x \ A) B" using assms by induct (auto simp: Set.filter_def) lemma image_fold_insert: assumes "finite A" shows "image f A = fold (\k A. Set.insert (f k) A) {} A" proof - interpret comp_fun_commute "\k A. Set.insert (f k) A" by standard auto show ?thesis using assms by (induct A) auto qed lemma Ball_fold: assumes "finite A" shows "Ball A P = fold (\k s. s \ P k) True A" proof - interpret comp_fun_commute "\k s. s \ P k" by standard auto show ?thesis using assms by (induct A) auto qed lemma Bex_fold: assumes "finite A" shows "Bex A P = fold (\k s. s \ P k) False A" proof - interpret comp_fun_commute "\k s. s \ P k" by standard auto show ?thesis using assms by (induct A) auto qed lemma comp_fun_commute_Pow_fold: "comp_fun_commute (\x A. A \ Set.insert x ` A)" by (clarsimp simp: fun_eq_iff comp_fun_commute_def) blast (* somewhat slow *) lemma Pow_fold: assumes "finite A" shows "Pow A = fold (\x A. A \ Set.insert x ` A) {{}} A" proof - interpret comp_fun_commute "\x A. A \ Set.insert x ` A" by (rule comp_fun_commute_Pow_fold) show ?thesis using assms by (induct A) (auto simp: Pow_insert) qed lemma fold_union_pair: assumes "finite B" shows "(\y\B. {(x, y)}) \ A = fold (\y. Set.insert (x, y)) A B" proof - interpret comp_fun_commute "\y. Set.insert (x, y)" by standard auto show ?thesis using assms by (induct arbitrary: A) simp_all qed lemma comp_fun_commute_product_fold: "finite B \ comp_fun_commute (\x z. fold (\y. Set.insert (x, y)) z B)" by standard (auto simp: fold_union_pair [symmetric]) lemma product_fold: assumes "finite A" "finite B" shows "A \ B = fold (\x z. fold (\y. Set.insert (x, y)) z B) {} A" using assms unfolding Sigma_def by (induct A) (simp_all add: comp_fun_commute.fold_insert[OF comp_fun_commute_product_fold] fold_union_pair) context complete_lattice begin lemma inf_Inf_fold_inf: assumes "finite A" shows "inf (Inf A) B = fold inf B A" proof - interpret comp_fun_idem inf by (fact comp_fun_idem_inf) from \finite A\ fold_fun_left_comm show ?thesis by (induct A arbitrary: B) (simp_all add: inf_commute fun_eq_iff) qed lemma sup_Sup_fold_sup: assumes "finite A" shows "sup (Sup A) B = fold sup B A" proof - interpret comp_fun_idem sup by (fact comp_fun_idem_sup) from \finite A\ fold_fun_left_comm show ?thesis by (induct A arbitrary: B) (simp_all add: sup_commute fun_eq_iff) qed lemma Inf_fold_inf: "finite A \ Inf A = fold inf top A" using inf_Inf_fold_inf [of A top] by (simp add: inf_absorb2) lemma Sup_fold_sup: "finite A \ Sup A = fold sup bot A" using sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2) lemma inf_INF_fold_inf: assumes "finite A" shows "inf B (\(f ` A)) = fold (inf \ f) B A" (is "?inf = ?fold") proof - interpret comp_fun_idem inf by (fact comp_fun_idem_inf) interpret comp_fun_idem "inf \ f" by (fact comp_comp_fun_idem) from \finite A\ have "?fold = ?inf" by (induct A arbitrary: B) (simp_all add: inf_left_commute) then show ?thesis .. qed lemma sup_SUP_fold_sup: assumes "finite A" shows "sup B (\(f ` A)) = fold (sup \ f) B A" (is "?sup = ?fold") proof - interpret comp_fun_idem sup by (fact comp_fun_idem_sup) interpret comp_fun_idem "sup \ f" by (fact comp_comp_fun_idem) from \finite A\ have "?fold = ?sup" by (induct A arbitrary: B) (simp_all add: sup_left_commute) then show ?thesis .. qed lemma INF_fold_inf: "finite A \ \(f ` A) = fold (inf \ f) top A" using inf_INF_fold_inf [of A top] by simp lemma SUP_fold_sup: "finite A \ \(f ` A) = fold (sup \ f) bot A" using sup_SUP_fold_sup [of A bot] by simp +lemma finite_Inf_in: + assumes "finite A" "A\{}" and inf: "\x y. \x \ A; y \ A\ \ inf x y \ A" + shows "Inf A \ A" +proof - + have "Inf B \ A" if "B \ A" "B\{}" for B + using finite_subset [OF \B \ A\ \finite A\] that + by (induction B) (use inf in \force+\) + then show ?thesis + by (simp add: assms) +qed + +lemma finite_Sup_in: + assumes "finite A" "A\{}" and sup: "\x y. \x \ A; y \ A\ \ sup x y \ A" + shows "Sup A \ A" +proof - + have "Sup B \ A" if "B \ A" "B\{}" for B + using finite_subset [OF \B \ A\ \finite A\] that + by (induction B) (use sup in \force+\) + then show ?thesis + by (simp add: assms) +qed + end subsection \Locales as mini-packages for fold operations\ subsubsection \The natural case\ locale folding = fixes f :: "'a \ 'b \ 'b" and z :: "'b" assumes comp_fun_commute: "f y \ f x = f x \ f y" begin interpretation fold?: comp_fun_commute f by standard (use comp_fun_commute in \simp add: fun_eq_iff\) definition F :: "'a set \ 'b" where eq_fold: "F A = fold f z A" lemma empty [simp]:"F {} = z" by (simp add: eq_fold) lemma infinite [simp]: "\ finite A \ F A = z" by (simp add: eq_fold) lemma insert [simp]: assumes "finite A" and "x \ A" shows "F (insert x A) = f x (F A)" proof - from fold_insert assms have "fold f z (insert x A) = f x (fold f z A)" by simp with \finite A\ show ?thesis by (simp add: eq_fold fun_eq_iff) qed lemma remove: assumes "finite A" and "x \ A" shows "F A = f x (F (A - {x}))" proof - from \x \ A\ obtain B where A: "A = insert x B" and "x \ B" by (auto dest: mk_disjoint_insert) moreover from \finite A\ A have "finite B" by simp ultimately show ?thesis by simp qed lemma insert_remove: "finite A \ F (insert x A) = f x (F (A - {x}))" by (cases "x \ A") (simp_all add: remove insert_absorb) end subsubsection \With idempotency\ locale folding_idem = folding + assumes comp_fun_idem: "f x \ f x = f x" begin declare insert [simp del] interpretation fold?: comp_fun_idem f by standard (insert comp_fun_commute comp_fun_idem, simp add: fun_eq_iff) lemma insert_idem [simp]: assumes "finite A" shows "F (insert x A) = f x (F A)" proof - from fold_insert_idem assms have "fold f z (insert x A) = f x (fold f z A)" by simp with \finite A\ show ?thesis by (simp add: eq_fold fun_eq_iff) qed end subsection \Finite cardinality\ text \ The traditional definition \<^prop>\card A \ LEAST n. \f. A = {f i |i. i < n}\ is ugly to work with. But now that we have \<^const>\fold\ things are easy: \ global_interpretation card: folding "\_. Suc" 0 defines card = "folding.F (\_. Suc) 0" by standard rule lemma card_infinite: "\ finite A \ card A = 0" by (fact card.infinite) lemma card_empty: "card {} = 0" by (fact card.empty) lemma card_insert_disjoint: "finite A \ x \ A \ card (insert x A) = Suc (card A)" by (fact card.insert) lemma card_insert_if: "finite A \ card (insert x A) = (if x \ A then card A else Suc (card A))" by auto (simp add: card.insert_remove card.remove) lemma card_ge_0_finite: "card A > 0 \ finite A" by (rule ccontr) simp lemma card_0_eq [simp]: "finite A \ card A = 0 \ A = {}" by (auto dest: mk_disjoint_insert) lemma finite_UNIV_card_ge_0: "finite (UNIV :: 'a set) \ card (UNIV :: 'a set) > 0" by (rule ccontr) simp lemma card_eq_0_iff: "card A = 0 \ A = {} \ \ finite A" by auto lemma card_range_greater_zero: "finite (range f) \ card (range f) > 0" by (rule ccontr) (simp add: card_eq_0_iff) lemma card_gt_0_iff: "0 < card A \ A \ {} \ finite A" by (simp add: neq0_conv [symmetric] card_eq_0_iff) lemma card_Suc_Diff1: "finite A \ x \ A \ Suc (card (A - {x})) = card A" apply (rule insert_Diff [THEN subst, where t = A]) apply assumption apply (simp del: insert_Diff_single) done lemma card_insert_le_m1: "n > 0 \ card y \ n - 1 \ card (insert x y) \ n" apply (cases "finite y") apply (cases "x \ y") apply (auto simp: insert_absorb) done lemma card_Diff_singleton: "finite A \ x \ A \ card (A - {x}) = card A - 1" by (simp add: card_Suc_Diff1 [symmetric]) lemma card_Diff_singleton_if: "finite A \ card (A - {x}) = (if x \ A then card A - 1 else card A)" by (simp add: card_Diff_singleton) lemma card_Diff_insert[simp]: assumes "finite A" and "a \ A" and "a \ B" shows "card (A - insert a B) = card (A - B) - 1" proof - have "A - insert a B = (A - B) - {a}" using assms by blast then show ?thesis using assms by (simp add: card_Diff_singleton) qed lemma card_insert: "finite A \ card (insert x A) = Suc (card (A - {x}))" by (fact card.insert_remove) lemma card_insert_le: "finite A \ card A \ card (insert x A)" by (simp add: card_insert_if) lemma card_Collect_less_nat[simp]: "card {i::nat. i < n} = n" by (induct n) (simp_all add:less_Suc_eq Collect_disj_eq) lemma card_Collect_le_nat[simp]: "card {i::nat. i \ n} = Suc n" using card_Collect_less_nat[of "Suc n"] by (simp add: less_Suc_eq_le) lemma card_mono: assumes "finite B" and "A \ B" shows "card A \ card B" proof - from assms have "finite A" by (auto intro: finite_subset) then show ?thesis using assms proof (induct A arbitrary: B) case empty then show ?case by simp next case (insert x A) then have "x \ B" by simp from insert have "A \ B - {x}" and "finite (B - {x})" by auto with insert.hyps have "card A \ card (B - {x})" by auto with \finite A\ \x \ A\ \finite B\ \x \ B\ show ?case by simp (simp only: card.remove) qed qed lemma card_seteq: "finite B \ (\A. A \ B \ card B \ card A \ A = B)" apply (induct rule: finite_induct) apply simp apply clarify apply (subgoal_tac "finite A \ A - {x} \ F") prefer 2 apply (blast intro: finite_subset, atomize) apply (drule_tac x = "A - {x}" in spec) apply (simp add: card_Diff_singleton_if split: if_split_asm) apply (case_tac "card A", auto) done lemma psubset_card_mono: "finite B \ A < B \ card A < card B" apply (simp add: psubset_eq linorder_not_le [symmetric]) apply (blast dest: card_seteq) done lemma card_Un_Int: assumes "finite A" "finite B" shows "card A + card B = card (A \ B) + card (A \ B)" using assms proof (induct A) case empty then show ?case by simp next case insert then show ?case by (auto simp add: insert_absorb Int_insert_left) qed lemma card_Un_disjoint: "finite A \ finite B \ A \ B = {} \ card (A \ B) = card A + card B" using card_Un_Int [of A B] by simp lemma card_Un_disjnt: "\finite A; finite B; disjnt A B\ \ card (A \ B) = card A + card B" by (simp add: card_Un_disjoint disjnt_def) lemma card_Un_le: "card (A \ B) \ card A + card B" proof (cases "finite A \ finite B") case True then show ?thesis using le_iff_add card_Un_Int [of A B] by auto qed auto lemma card_Diff_subset: assumes "finite B" and "B \ A" shows "card (A - B) = card A - card B" using assms proof (cases "finite A") case False with assms show ?thesis by simp next case True with assms show ?thesis by (induct B arbitrary: A) simp_all qed lemma card_Diff_subset_Int: assumes "finite (A \ B)" shows "card (A - B) = card A - card (A \ B)" proof - have "A - B = A - A \ B" by auto with assms show ?thesis by (simp add: card_Diff_subset) qed lemma diff_card_le_card_Diff: assumes "finite B" shows "card A - card B \ card (A - B)" proof - have "card A - card B \ card A - card (A \ B)" using card_mono[OF assms Int_lower2, of A] by arith also have "\ = card (A - B)" using assms by (simp add: card_Diff_subset_Int) finally show ?thesis . qed lemma card_le_sym_Diff: assumes "finite A" "finite B" "card A \ card B" shows "card(A - B) \ card(B - A)" proof - have "card(A - B) = card A - card (A \ B)" using assms(1,2) by(simp add: card_Diff_subset_Int) also have "\ \ card B - card (A \ B)" using assms(3) by linarith also have "\ = card(B - A)" using assms(1,2) by(simp add: card_Diff_subset_Int Int_commute) finally show ?thesis . qed lemma card_less_sym_Diff: assumes "finite A" "finite B" "card A < card B" shows "card(A - B) < card(B - A)" proof - have "card(A - B) = card A - card (A \ B)" using assms(1,2) by(simp add: card_Diff_subset_Int) also have "\ < card B - card (A \ B)" using assms(1,3) by (simp add: card_mono diff_less_mono) also have "\ = card(B - A)" using assms(1,2) by(simp add: card_Diff_subset_Int Int_commute) finally show ?thesis . qed lemma card_Diff1_less: "finite A \ x \ A \ card (A - {x}) < card A" by (rule Suc_less_SucD) (simp add: card_Suc_Diff1 del: card_Diff_insert) lemma card_Diff2_less: "finite A \ x \ A \ y \ A \ card (A - {x} - {y}) < card A" apply (cases "x = y") apply (simp add: card_Diff1_less del:card_Diff_insert) apply (rule less_trans) prefer 2 apply (auto intro!: card_Diff1_less simp del: card_Diff_insert) done lemma card_Diff1_le: "finite A \ card (A - {x}) \ card A" by (cases "x \ A") (simp_all add: card_Diff1_less less_imp_le) lemma card_psubset: "finite B \ A \ B \ card A < card B \ A < B" by (erule psubsetI) blast lemma card_le_inj: assumes fA: "finite A" and fB: "finite B" and c: "card A \ card B" shows "\f. f ` A \ B \ inj_on f A" using fA fB c proof (induct arbitrary: B rule: finite_induct) case empty then show ?case by simp next case (insert x s t) then show ?case proof (induct rule: finite_induct [OF insert.prems(1)]) case 1 then show ?case by simp next case (2 y t) from "2.prems"(1,2,5) "2.hyps"(1,2) have cst: "card s \ card t" by simp from "2.prems"(3) [OF "2.hyps"(1) cst] obtain f where "f ` s \ t" "inj_on f s" by blast with "2.prems"(2) "2.hyps"(2) show ?case apply - apply (rule exI[where x = "\z. if z = x then y else f z"]) apply (auto simp add: inj_on_def) done qed qed lemma card_subset_eq: assumes fB: "finite B" and AB: "A \ B" and c: "card A = card B" shows "A = B" proof - from fB AB have fA: "finite A" by (auto intro: finite_subset) from fA fB have fBA: "finite (B - A)" by auto have e: "A \ (B - A) = {}" by blast have eq: "A \ (B - A) = B" using AB by blast from card_Un_disjoint[OF fA fBA e, unfolded eq c] have "card (B - A) = 0" by arith then have "B - A = {}" unfolding card_eq_0_iff using fA fB by simp with AB show "A = B" by blast qed lemma insert_partition: "x \ F \ \c1 \ insert x F. \c2 \ insert x F. c1 \ c2 \ c1 \ c2 = {} \ x \ \F = {}" by auto (* somewhat slow *) lemma finite_psubset_induct [consumes 1, case_names psubset]: assumes finite: "finite A" and major: "\A. finite A \ (\B. B \ A \ P B) \ P A" shows "P A" using finite proof (induct A taking: card rule: measure_induct_rule) case (less A) have fin: "finite A" by fact have ih: "card B < card A \ finite B \ P B" for B by fact have "P B" if "B \ A" for B proof - from that have "card B < card A" using psubset_card_mono fin by blast moreover from that have "B \ A" by auto then have "finite B" using fin finite_subset by blast ultimately show ?thesis using ih by simp qed with fin show "P A" using major by blast qed lemma finite_induct_select [consumes 1, case_names empty select]: assumes "finite S" and "P {}" and select: "\T. T \ S \ P T \ \s\S - T. P (insert s T)" shows "P S" proof - have "0 \ card S" by simp then have "\T \ S. card T = card S \ P T" proof (induct rule: dec_induct) case base with \P {}\ show ?case by (intro exI[of _ "{}"]) auto next case (step n) then obtain T where T: "T \ S" "card T = n" "P T" by auto with \n < card S\ have "T \ S" "P T" by auto with select[of T] obtain s where "s \ S" "s \ T" "P (insert s T)" by auto with step(2) T \finite S\ show ?case by (intro exI[of _ "insert s T"]) (auto dest: finite_subset) qed with \finite S\ show "P S" by (auto dest: card_subset_eq) qed lemma remove_induct [case_names empty infinite remove]: assumes empty: "P ({} :: 'a set)" and infinite: "\ finite B \ P B" and remove: "\A. finite A \ A \ {} \ A \ B \ (\x. x \ A \ P (A - {x})) \ P A" shows "P B" proof (cases "finite B") case False then show ?thesis by (rule infinite) next case True define A where "A = B" with True have "finite A" "A \ B" by simp_all then show "P A" proof (induct "card A" arbitrary: A) case 0 then have "A = {}" by auto with empty show ?case by simp next case (Suc n A) from \A \ B\ and \finite B\ have "finite A" by (rule finite_subset) moreover from Suc.hyps have "A \ {}" by auto moreover note \A \ B\ moreover have "P (A - {x})" if x: "x \ A" for x using x Suc.prems \Suc n = card A\ by (intro Suc) auto ultimately show ?case by (rule remove) qed qed lemma finite_remove_induct [consumes 1, case_names empty remove]: fixes P :: "'a set \ bool" assumes "finite B" and "P {}" and "\A. finite A \ A \ {} \ A \ B \ (\x. x \ A \ P (A - {x})) \ P A" defines "B' \ B" shows "P B'" by (induct B' rule: remove_induct) (simp_all add: assms) text \Main cardinality theorem.\ lemma card_partition [rule_format]: "finite C \ finite (\C) \ (\c\C. card c = k) \ (\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {}) \ k * card C = card (\C)" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) then show ?case by (simp add: card_Un_disjoint insert_partition finite_subset [of _ "\(insert _ _)"]) qed lemma card_eq_UNIV_imp_eq_UNIV: assumes fin: "finite (UNIV :: 'a set)" and card: "card A = card (UNIV :: 'a set)" shows "A = (UNIV :: 'a set)" proof show "A \ UNIV" by simp show "UNIV \ A" proof show "x \ A" for x proof (rule ccontr) assume "x \ A" then have "A \ UNIV" by auto with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono) with card show False by simp qed qed qed text \The form of a finite set of given cardinality\ lemma card_eq_SucD: assumes "card A = Suc k" shows "\b B. A = insert b B \ b \ B \ card B = k \ (k = 0 \ B = {})" proof - have fin: "finite A" using assms by (auto intro: ccontr) moreover have "card A \ 0" using assms by auto ultimately obtain b where b: "b \ A" by auto show ?thesis proof (intro exI conjI) show "A = insert b (A - {b})" using b by blast show "b \ A - {b}" by blast show "card (A - {b}) = k" and "k = 0 \ A - {b} = {}" using assms b fin by (fastforce dest: mk_disjoint_insert)+ qed qed lemma card_Suc_eq: "card A = Suc k \ (\b B. A = insert b B \ b \ B \ card B = k \ (k = 0 \ B = {}))" apply (auto elim!: card_eq_SucD) apply (subst card.insert) apply (auto simp add: intro:ccontr) done lemma card_1_singletonE: assumes "card A = 1" obtains x where "A = {x}" using assms by (auto simp: card_Suc_eq) lemma is_singleton_altdef: "is_singleton A \ card A = 1" unfolding is_singleton_def by (auto elim!: card_1_singletonE is_singletonE simp del: One_nat_def) lemma card_1_singleton_iff: "card A = Suc 0 \ (\x. A = {x})" by (simp add: card_Suc_eq) lemma card_le_Suc0_iff_eq: assumes "finite A" shows "card A \ Suc 0 \ (\a1 \ A. \a2 \ A. a1 = a2)" (is "?C = ?A") proof assume ?C thus ?A using assms by (auto simp: le_Suc_eq dest: card_eq_SucD) next assume ?A show ?C proof cases assume "A = {}" thus ?C using \?A\ by simp next assume "A \ {}" then obtain a where "A = {a}" using \?A\ by blast thus ?C by simp qed qed lemma card_le_Suc_iff: "Suc n \ card A = (\a B. A = insert a B \ a \ B \ n \ card B \ finite B)" apply(cases "finite A") apply (fastforce simp: card_Suc_eq less_eq_nat.simps(2) insert_eq_iff dest: subset_singletonD split: nat.splits if_splits) by auto lemma finite_fun_UNIVD2: assumes fin: "finite (UNIV :: ('a \ 'b) set)" shows "finite (UNIV :: 'b set)" proof - from fin have "finite (range (\f :: 'a \ 'b. f arbitrary))" for arbitrary by (rule finite_imageI) moreover have "UNIV = range (\f :: 'a \ 'b. f arbitrary)" for arbitrary by (rule UNIV_eq_I) auto ultimately show "finite (UNIV :: 'b set)" by simp qed lemma card_UNIV_unit [simp]: "card (UNIV :: unit set) = 1" unfolding UNIV_unit by simp lemma infinite_arbitrarily_large: assumes "\ finite A" shows "\B. finite B \ card B = n \ B \ A" proof (induction n) case 0 show ?case by (intro exI[of _ "{}"]) auto next case (Suc n) then obtain B where B: "finite B \ card B = n \ B \ A" .. with \\ finite A\ have "A \ B" by auto with B have "B \ A" by auto then have "\x. x \ A - B" by (elim psubset_imp_ex_mem) then obtain x where x: "x \ A - B" .. with B have "finite (insert x B) \ card (insert x B) = Suc n \ insert x B \ A" by auto then show "\B. finite B \ card B = Suc n \ B \ A" .. qed text \Sometimes, to prove that a set is finite, it is convenient to work with finite subsets and to show that their cardinalities are uniformly bounded. This possibility is formalized in the next criterion.\ lemma finite_if_finite_subsets_card_bdd: assumes "\G. G \ F \ finite G \ card G \ C" shows "finite F \ card F \ C" proof (cases "finite F") case False obtain n::nat where n: "n > max C 0" by auto obtain G where G: "G \ F" "card G = n" using infinite_arbitrarily_large[OF False] by auto hence "finite G" using \n > max C 0\ using card_infinite gr_implies_not0 by blast hence False using assms G n not_less by auto thus ?thesis .. next case True thus ?thesis using assms[of F] by auto qed subsubsection \Cardinality of image\ lemma card_image_le: "finite A \ card (f ` A) \ card A" by (induct rule: finite_induct) (simp_all add: le_SucI card_insert_if) lemma card_image: "inj_on f A \ card (f ` A) = card A" proof (induct A rule: infinite_finite_induct) case (infinite A) then have "\ finite (f ` A)" by (auto dest: finite_imageD) with infinite show ?case by simp qed simp_all lemma bij_betw_same_card: "bij_betw f A B \ card A = card B" by (auto simp: card_image bij_betw_def) lemma endo_inj_surj: "finite A \ f ` A \ A \ inj_on f A \ f ` A = A" by (simp add: card_seteq card_image) lemma eq_card_imp_inj_on: assumes "finite A" "card(f ` A) = card A" shows "inj_on f A" using assms proof (induct rule:finite_induct) case empty show ?case by simp next case (insert x A) then show ?case using card_image_le [of A f] by (simp add: card_insert_if split: if_splits) qed lemma inj_on_iff_eq_card: "finite A \ inj_on f A \ card (f ` A) = card A" by (blast intro: card_image eq_card_imp_inj_on) lemma card_inj_on_le: assumes "inj_on f A" "f ` A \ B" "finite B" shows "card A \ card B" proof - have "finite A" using assms by (blast intro: finite_imageD dest: finite_subset) then show ?thesis using assms by (force intro: card_mono simp: card_image [symmetric]) qed lemma inj_on_iff_card_le: "\ finite A; finite B \ \ (\f. inj_on f A \ f ` A \ B) = (card A \ card B)" using card_inj_on_le[of _ A B] card_le_inj[of A B] by blast lemma surj_card_le: "finite A \ B \ f ` A \ card B \ card A" by (blast intro: card_image_le card_mono le_trans) lemma card_bij_eq: "inj_on f A \ f ` A \ B \ inj_on g B \ g ` B \ A \ finite A \ finite B \ card A = card B" by (auto intro: le_antisym card_inj_on_le) lemma bij_betw_finite: "bij_betw f A B \ finite A \ finite B" unfolding bij_betw_def using finite_imageD [of f A] by auto lemma inj_on_finite: "inj_on f A \ f ` A \ B \ finite B \ finite A" using finite_imageD finite_subset by blast lemma card_vimage_inj: "inj f \ A \ range f \ card (f -` A) = card A" by (auto 4 3 simp: subset_image_iff inj_vimage_image_eq intro: card_image[symmetric, OF subset_inj_on]) subsubsection \Pigeonhole Principles\ lemma pigeonhole: "card A > card (f ` A) \ \ inj_on f A " by (auto dest: card_image less_irrefl_nat) lemma pigeonhole_infinite: assumes "\ finite A" and "finite (f`A)" shows "\a0\A. \ finite {a\A. f a = f a0}" using assms(2,1) proof (induct "f`A" arbitrary: A rule: finite_induct) case empty then show ?case by simp next case (insert b F) show ?case proof (cases "finite {a\A. f a = b}") case True with \\ finite A\ have "\ finite (A - {a\A. f a = b})" by simp also have "A - {a\A. f a = b} = {a\A. f a \ b}" by blast finally have "\ finite {a\A. f a \ b}" . from insert(3)[OF _ this] insert(2,4) show ?thesis by simp (blast intro: rev_finite_subset) next case False then have "{a \ A. f a = b} \ {}" by force with False show ?thesis by blast qed qed lemma pigeonhole_infinite_rel: assumes "\ finite A" and "finite B" and "\a\A. \b\B. R a b" shows "\b\B. \ finite {a:A. R a b}" proof - let ?F = "\a. {b\B. R a b}" from finite_Pow_iff[THEN iffD2, OF \finite B\] have "finite (?F ` A)" by (blast intro: rev_finite_subset) from pigeonhole_infinite [where f = ?F, OF assms(1) this] obtain a0 where "a0 \ A" and infinite: "\ finite {a\A. ?F a = ?F a0}" .. obtain b0 where "b0 \ B" and "R a0 b0" using \a0 \ A\ assms(3) by blast have "finite {a\A. ?F a = ?F a0}" if "finite {a\A. R a b0}" using \b0 \ B\ \R a0 b0\ that by (blast intro: rev_finite_subset) with infinite \b0 \ B\ show ?thesis by blast qed subsubsection \Cardinality of sums\ lemma card_Plus: assumes "finite A" "finite B" shows "card (A <+> B) = card A + card B" proof - have "Inl`A \ Inr`B = {}" by fast with assms show ?thesis by (simp add: Plus_def card_Un_disjoint card_image) qed lemma card_Plus_conv_if: "card (A <+> B) = (if finite A \ finite B then card A + card B else 0)" by (auto simp add: card_Plus) text \Relates to equivalence classes. Based on a theorem of F. Kammüller.\ lemma dvd_partition: assumes f: "finite (\C)" and "\c\C. k dvd card c" "\c1\C. \c2\C. c1 \ c2 \ c1 \ c2 = {}" shows "k dvd card (\C)" proof - have "finite C" by (rule finite_UnionD [OF f]) then show ?thesis using assms proof (induct rule: finite_induct) case empty show ?case by simp next case insert then show ?case apply simp apply (subst card_Un_disjoint) apply (auto simp add: disjoint_eq_subset_Compl) done qed qed subsubsection \Relating injectivity and surjectivity\ lemma finite_surj_inj: assumes "finite A" "A \ f ` A" shows "inj_on f A" proof - have "f ` A = A" by (rule card_seteq [THEN sym]) (auto simp add: assms card_image_le) then show ?thesis using assms by (simp add: eq_card_imp_inj_on) qed lemma finite_UNIV_surj_inj: "finite(UNIV:: 'a set) \ surj f \ inj f" for f :: "'a \ 'a" by (blast intro: finite_surj_inj subset_UNIV) lemma finite_UNIV_inj_surj: "finite(UNIV:: 'a set) \ inj f \ surj f" for f :: "'a \ 'a" by (fastforce simp:surj_def dest!: endo_inj_surj) lemma surjective_iff_injective_gen: assumes fS: "finite S" and fT: "finite T" and c: "card S = card T" and ST: "f ` S \ T" shows "(\y \ T. \x \ S. f x = y) \ inj_on f S" (is "?lhs \ ?rhs") proof assume h: "?lhs" { fix x y assume x: "x \ S" assume y: "y \ S" assume f: "f x = f y" from x fS have S0: "card S \ 0" by auto have "x = y" proof (rule ccontr) assume xy: "\ ?thesis" have th: "card S \ card (f ` (S - {y}))" unfolding c proof (rule card_mono) show "finite (f ` (S - {y}))" by (simp add: fS) have "\x \ y; x \ S; z \ S; f x = f y\ \ \x \ S. x \ y \ f z = f x" for z by (case_tac "z = y \ z = x") auto then show "T \ f ` (S - {y})" using h xy x y f by fastforce qed also have " \ \ card (S - {y})" by (simp add: card_image_le fS) also have "\ \ card S - 1" using y fS by simp finally show False using S0 by arith qed } then show ?rhs unfolding inj_on_def by blast next assume h: ?rhs have "f ` S = T" by (simp add: ST c card_image card_subset_eq fT h) then show ?lhs by blast qed hide_const (open) Finite_Set.fold subsection \Infinite Sets\ text \ Some elementary facts about infinite sets, mostly by Stephan Merz. Beware! Because "infinite" merely abbreviates a negation, these lemmas may not work well with \blast\. \ abbreviation infinite :: "'a set \ bool" where "infinite S \ \ finite S" text \ Infinite sets are non-empty, and if we remove some elements from an infinite set, the result is still infinite. \ lemma infinite_UNIV_nat [iff]: "infinite (UNIV :: nat set)" proof assume "finite (UNIV :: nat set)" with finite_UNIV_inj_surj [of Suc] show False by simp (blast dest: Suc_neq_Zero surjD) qed lemma infinite_UNIV_char_0: "infinite (UNIV :: 'a::semiring_char_0 set)" proof assume "finite (UNIV :: 'a set)" with subset_UNIV have "finite (range of_nat :: 'a set)" by (rule finite_subset) moreover have "inj (of_nat :: nat \ 'a)" by (simp add: inj_on_def) ultimately have "finite (UNIV :: nat set)" by (rule finite_imageD) then show False by simp qed lemma infinite_imp_nonempty: "infinite S \ S \ {}" by auto lemma infinite_remove: "infinite S \ infinite (S - {a})" by simp lemma Diff_infinite_finite: assumes "finite T" "infinite S" shows "infinite (S - T)" using \finite T\ proof induct from \infinite S\ show "infinite (S - {})" by auto next fix T x assume ih: "infinite (S - T)" have "S - (insert x T) = (S - T) - {x}" by (rule Diff_insert) with ih show "infinite (S - (insert x T))" by (simp add: infinite_remove) qed lemma Un_infinite: "infinite S \ infinite (S \ T)" by simp lemma infinite_Un: "infinite (S \ T) \ infinite S \ infinite T" by simp lemma infinite_super: assumes "S \ T" and "infinite S" shows "infinite T" proof assume "finite T" with \S \ T\ have "finite S" by (simp add: finite_subset) with \infinite S\ show False by simp qed proposition infinite_coinduct [consumes 1, case_names infinite]: assumes "X A" and step: "\A. X A \ \x\A. X (A - {x}) \ infinite (A - {x})" shows "infinite A" proof assume "finite A" then show False using \X A\ proof (induction rule: finite_psubset_induct) case (psubset A) then obtain x where "x \ A" "X (A - {x}) \ infinite (A - {x})" using local.step psubset.prems by blast then have "X (A - {x})" using psubset.hyps by blast show False apply (rule psubset.IH [where B = "A - {x}"]) apply (use \x \ A\ in blast) apply (simp add: \X (A - {x})\) done qed qed text \ For any function with infinite domain and finite range there is some element that is the image of infinitely many domain elements. In particular, any infinite sequence of elements from a finite set contains some element that occurs infinitely often. \ lemma inf_img_fin_dom': assumes img: "finite (f ` A)" and dom: "infinite A" shows "\y \ f ` A. infinite (f -` {y} \ A)" proof (rule ccontr) have "A \ (\y\f ` A. f -` {y} \ A)" by auto moreover assume "\ ?thesis" with img have "finite (\y\f ` A. f -` {y} \ A)" by blast ultimately have "finite A" by (rule finite_subset) with dom show False by contradiction qed lemma inf_img_fin_domE': assumes "finite (f ` A)" and "infinite A" obtains y where "y \ f`A" and "infinite (f -` {y} \ A)" using assms by (blast dest: inf_img_fin_dom') lemma inf_img_fin_dom: assumes img: "finite (f`A)" and dom: "infinite A" shows "\y \ f`A. infinite (f -` {y})" using inf_img_fin_dom'[OF assms] by auto lemma inf_img_fin_domE: assumes "finite (f`A)" and "infinite A" obtains y where "y \ f`A" and "infinite (f -` {y})" using assms by (blast dest: inf_img_fin_dom) proposition finite_image_absD: "finite (abs ` S) \ finite S" for S :: "'a::linordered_ring set" by (rule ccontr) (auto simp: abs_eq_iff vimage_def dest: inf_img_fin_dom) subsection \The finite powerset operator\ definition Fpow :: "'a set \ 'a set set" where "Fpow A \ {X. X \ A \ finite X}" lemma Fpow_mono: "A \ B \ Fpow A \ Fpow B" unfolding Fpow_def by auto lemma empty_in_Fpow: "{} \ Fpow A" unfolding Fpow_def by auto lemma Fpow_not_empty: "Fpow A \ {}" using empty_in_Fpow by blast lemma Fpow_subset_Pow: "Fpow A \ Pow A" unfolding Fpow_def by auto lemma Fpow_Pow_finite: "Fpow A = Pow A Int {A. finite A}" unfolding Fpow_def Pow_def by blast lemma inj_on_image_Fpow: assumes "inj_on f A" shows "inj_on (image f) (Fpow A)" using assms Fpow_subset_Pow[of A] subset_inj_on[of "image f" "Pow A"] inj_on_image_Pow by blast lemma image_Fpow_mono: assumes "f ` A \ B" shows "(image f) ` (Fpow A) \ Fpow B" using assms by(unfold Fpow_def, auto) end diff --git a/src/HOL/Fun_Def.thy b/src/HOL/Fun_Def.thy --- a/src/HOL/Fun_Def.thy +++ b/src/HOL/Fun_Def.thy @@ -1,307 +1,310 @@ (* Title: HOL/Fun_Def.thy Author: Alexander Krauss, TU Muenchen *) section \Function Definitions and Termination Proofs\ theory Fun_Def imports Basic_BNF_LFPs Partial_Function SAT keywords "function" "termination" :: thy_goal_defn and "fun" "fun_cases" :: thy_defn begin subsection \Definitions with default value\ definition THE_default :: "'a \ ('a \ bool) \ 'a" where "THE_default d P = (if (\!x. P x) then (THE x. P x) else d)" lemma THE_defaultI': "\!x. P x \ P (THE_default d P)" by (simp add: theI' THE_default_def) lemma THE_default1_equality: "\!x. P x \ P a \ THE_default d P = a" by (simp add: the1_equality THE_default_def) lemma THE_default_none: "\ (\!x. P x) \ THE_default d P = d" by (simp add: THE_default_def) lemma fundef_ex1_existence: assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" shows "G x (f x)" apply (simp only: f_def) apply (rule THE_defaultI') apply (rule ex1) done lemma fundef_ex1_uniqueness: assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" assumes elm: "G x (h x)" shows "h x = f x" apply (simp only: f_def) apply (rule THE_default1_equality [symmetric]) apply (rule ex1) apply (rule elm) done lemma fundef_ex1_iff: assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" shows "(G x y) = (f x = y)" apply (auto simp:ex1 f_def THE_default1_equality) apply (rule THE_defaultI') apply (rule ex1) done lemma fundef_default_value: assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes graph: "\x y. G x y \ D x" assumes "\ D x" shows "f x = d x" proof - have "\(\y. G x y)" proof assume "\y. G x y" then have "D x" using graph .. with \\ D x\ show False .. qed then have "\(\!y. G x y)" by blast then show ?thesis unfolding f_def by (rule THE_default_none) qed definition in_rel_def[simp]: "in_rel R x y \ (x, y) \ R" lemma wf_in_rel: "wf R \ wfP (in_rel R)" by (simp add: wfP_def) ML_file \Tools/Function/function_core.ML\ ML_file \Tools/Function/mutual.ML\ ML_file \Tools/Function/pattern_split.ML\ ML_file \Tools/Function/relation.ML\ ML_file \Tools/Function/function_elims.ML\ method_setup relation = \ Args.term >> (fn t => fn ctxt => SIMPLE_METHOD' (Function_Relation.relation_infer_tac ctxt t)) \ "prove termination using a user-specified wellfounded relation" ML_file \Tools/Function/function.ML\ ML_file \Tools/Function/pat_completeness.ML\ method_setup pat_completeness = \ Scan.succeed (SIMPLE_METHOD' o Pat_Completeness.pat_completeness_tac) \ "prove completeness of (co)datatype patterns" ML_file \Tools/Function/fun.ML\ ML_file \Tools/Function/induction_schema.ML\ method_setup induction_schema = \ Scan.succeed (CONTEXT_TACTIC oo Induction_Schema.induction_schema_tac) \ "prove an induction principle" subsection \Measure functions\ inductive is_measure :: "('a \ nat) \ bool" where is_measure_trivial: "is_measure f" named_theorems measure_function "rules that guide the heuristic generation of measure functions" ML_file \Tools/Function/measure_functions.ML\ lemma measure_size[measure_function]: "is_measure size" by (rule is_measure_trivial) lemma measure_fst[measure_function]: "is_measure f \ is_measure (\p. f (fst p))" by (rule is_measure_trivial) lemma measure_snd[measure_function]: "is_measure f \ is_measure (\p. f (snd p))" by (rule is_measure_trivial) ML_file \Tools/Function/lexicographic_order.ML\ method_setup lexicographic_order = \ Method.sections clasimp_modifiers >> (K (SIMPLE_METHOD o Lexicographic_Order.lexicographic_order_tac false)) \ "termination prover for lexicographic orderings" subsection \Congruence rules\ lemma let_cong [fundef_cong]: "M = N \ (\x. x = N \ f x = g x) \ Let M f = Let N g" unfolding Let_def by blast lemmas [fundef_cong] = if_cong image_cong bex_cong ball_cong imp_cong map_option_cong Option.bind_cong lemma split_cong [fundef_cong]: "(\x y. (x, y) = q \ f x y = g x y) \ p = q \ case_prod f p = case_prod g q" by (auto simp: split_def) lemma comp_cong [fundef_cong]: "f (g x) = f' (g' x') \ (f \ g) x = (f' \ g') x'" by (simp only: o_apply) subsection \Simp rules for termination proofs\ declare trans_less_add1[termination_simp] trans_less_add2[termination_simp] trans_le_add1[termination_simp] trans_le_add2[termination_simp] less_imp_le_nat[termination_simp] le_imp_less_Suc[termination_simp] lemma size_prod_simp[termination_simp]: "size_prod f g p = f (fst p) + g (snd p) + Suc 0" by (induct p) auto subsection \Decomposition\ lemma less_by_empty: "A = {} \ A \ B" and union_comp_emptyL: "A O C = {} \ B O C = {} \ (A \ B) O C = {}" and union_comp_emptyR: "A O B = {} \ A O C = {} \ A O (B \ C) = {}" and wf_no_loop: "R O R = {} \ wf R" by (auto simp add: wf_comp_self [of R]) subsection \Reduction pairs\ definition "reduction_pair P \ wf (fst P) \ fst P O snd P \ fst P" lemma reduction_pairI[intro]: "wf R \ R O S \ R \ reduction_pair (R, S)" by (auto simp: reduction_pair_def) lemma reduction_pair_lemma: assumes rp: "reduction_pair P" assumes "R \ fst P" assumes "S \ snd P" assumes "wf S" shows "wf (R \ S)" proof - from rp \S \ snd P\ have "wf (fst P)" "fst P O S \ fst P" unfolding reduction_pair_def by auto with \wf S\ have "wf (fst P \ S)" by (auto intro: wf_union_compatible) moreover from \R \ fst P\ have "R \ S \ fst P \ S" by auto ultimately show ?thesis by (rule wf_subset) qed definition "rp_inv_image = (\(R,S) f. (inv_image R f, inv_image S f))" lemma rp_inv_image_rp: "reduction_pair P \ reduction_pair (rp_inv_image P f)" unfolding reduction_pair_def rp_inv_image_def split_def by force subsection \Concrete orders for SCNP termination proofs\ definition "pair_less = less_than <*lex*> less_than" definition "pair_leq = pair_less\<^sup>=" definition "max_strict = max_ext pair_less" definition "max_weak = max_ext pair_leq \ {({}, {})}" definition "min_strict = min_ext pair_less" definition "min_weak = min_ext pair_leq \ {({}, {})}" lemma wf_pair_less[simp]: "wf pair_less" by (auto simp: pair_less_def) lemma total_pair_less [iff]: "total_on A pair_less" and trans_pair_less [iff]: "trans pair_less" by (auto simp: total_on_def pair_less_def) text \Introduction rules for \pair_less\/\pair_leq\\ lemma pair_leqI1: "a < b \ ((a, s), (b, t)) \ pair_leq" and pair_leqI2: "a \ b \ s \ t \ ((a, s), (b, t)) \ pair_leq" and pair_lessI1: "a < b \ ((a, s), (b, t)) \ pair_less" and pair_lessI2: "a \ b \ s < t \ ((a, s), (b, t)) \ pair_less" by (auto simp: pair_leq_def pair_less_def) +lemma pair_less_iff1 [simp]: "((x,y), (x,z)) \ pair_less \ yIntroduction rules for max\ lemma smax_emptyI: "finite Y \ Y \ {} \ ({}, Y) \ max_strict" and smax_insertI: "y \ Y \ (x, y) \ pair_less \ (X, Y) \ max_strict \ (insert x X, Y) \ max_strict" and wmax_emptyI: "finite X \ ({}, X) \ max_weak" and wmax_insertI: "y \ YS \ (x, y) \ pair_leq \ (XS, YS) \ max_weak \ (insert x XS, YS) \ max_weak" by (auto simp: max_strict_def max_weak_def elim!: max_ext.cases) text \Introduction rules for min\ lemma smin_emptyI: "X \ {} \ (X, {}) \ min_strict" and smin_insertI: "x \ XS \ (x, y) \ pair_less \ (XS, YS) \ min_strict \ (XS, insert y YS) \ min_strict" and wmin_emptyI: "(X, {}) \ min_weak" and wmin_insertI: "x \ XS \ (x, y) \ pair_leq \ (XS, YS) \ min_weak \ (XS, insert y YS) \ min_weak" by (auto simp: min_strict_def min_weak_def min_ext_def) text \Reduction Pairs.\ lemma max_ext_compat: assumes "R O S \ R" shows "max_ext R O (max_ext S \ {({}, {})}) \ max_ext R" using assms apply auto apply (elim max_ext.cases) apply rule apply auto[3] apply (drule_tac x=xa in meta_spec) apply simp apply (erule bexE) apply (drule_tac x=xb in meta_spec) apply auto done lemma max_rpair_set: "reduction_pair (max_strict, max_weak)" unfolding max_strict_def max_weak_def apply (intro reduction_pairI max_ext_wf) apply simp apply (rule max_ext_compat) apply (auto simp: pair_less_def pair_leq_def) done lemma min_ext_compat: assumes "R O S \ R" shows "min_ext R O (min_ext S \ {({},{})}) \ min_ext R" using assms apply (auto simp: min_ext_def) apply (drule_tac x=ya in bspec, assumption) apply (erule bexE) apply (drule_tac x=xc in bspec) apply assumption apply auto done lemma min_rpair_set: "reduction_pair (min_strict, min_weak)" unfolding min_strict_def min_weak_def apply (intro reduction_pairI min_ext_wf) apply simp apply (rule min_ext_compat) apply (auto simp: pair_less_def pair_leq_def) done subsection \Yet another induction principle on the natural numbers\ lemma nat_descend_induct [case_names base descend]: fixes P :: "nat \ bool" assumes H1: "\k. k > n \ P k" assumes H2: "\k. k \ n \ (\i. i > k \ P i) \ P k" shows "P m" using assms by induction_schema (force intro!: wf_measure [of "\k. Suc n - k"])+ subsection \Tool setup\ ML_file \Tools/Function/termination.ML\ ML_file \Tools/Function/scnp_solve.ML\ ML_file \Tools/Function/scnp_reconstruct.ML\ ML_file \Tools/Function/fun_cases.ML\ ML_val \ \setup inactive\ \ Context.theory_map (Function_Common.set_termination_prover (K (ScnpReconstruct.decomp_scnp_tac [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS]))) \ end diff --git a/src/HOL/Library/Infinite_Set.thy b/src/HOL/Library/Infinite_Set.thy --- a/src/HOL/Library/Infinite_Set.thy +++ b/src/HOL/Library/Infinite_Set.thy @@ -1,607 +1,614 @@ (* Title: HOL/Library/Infinite_Set.thy Author: Stephan Merz *) section \Infinite Sets and Related Concepts\ theory Infinite_Set imports Main begin subsection \The set of natural numbers is infinite\ lemma infinite_nat_iff_unbounded_le: "infinite S \ (\m. \n\m. n \ S)" for S :: "nat set" using frequently_cofinite[of "\x. x \ S"] by (simp add: cofinite_eq_sequentially frequently_def eventually_sequentially) lemma infinite_nat_iff_unbounded: "infinite S \ (\m. \n>m. n \ S)" for S :: "nat set" using frequently_cofinite[of "\x. x \ S"] by (simp add: cofinite_eq_sequentially frequently_def eventually_at_top_dense) lemma finite_nat_iff_bounded: "finite S \ (\k. S \ {.. (\k. S \ {.. k})" for S :: "nat set" using infinite_nat_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) lemma finite_nat_bounded: "finite S \ \k. S \ {.. For a set of natural numbers to be infinite, it is enough to know that for any number larger than some \k\, there is some larger number that is an element of the set. \ lemma unbounded_k_infinite: "\m>k. \n>m. n \ S \ infinite (S::nat set)" apply (clarsimp simp add: finite_nat_set_iff_bounded) apply (drule_tac x="Suc (max m k)" in spec) using less_Suc_eq apply fastforce done lemma nat_not_finite: "finite (UNIV::nat set) \ R" by simp lemma range_inj_infinite: fixes f :: "nat \ 'a" assumes "inj f" shows "infinite (range f)" proof assume "finite (range f)" from this assms have "finite (UNIV::nat set)" by (rule finite_imageD) then show False by simp qed subsection \The set of integers is also infinite\ lemma infinite_int_iff_infinite_nat_abs: "infinite S \ infinite ((nat \ abs) ` S)" for S :: "int set" proof (unfold Not_eq_iff, rule iffI) assume "finite ((nat \ abs) ` S)" then have "finite (nat ` (abs ` S))" by (simp add: image_image cong: image_cong) moreover have "inj_on nat (abs ` S)" by (rule inj_onI) auto ultimately have "finite (abs ` S)" by (rule finite_imageD) then show "finite S" by (rule finite_image_absD) qed simp proposition infinite_int_iff_unbounded_le: "infinite S \ (\m. \n. \n\ \ m \ n \ S)" for S :: "int set" by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded_le o_def image_def) (metis abs_ge_zero nat_le_eq_zle le_nat_iff) proposition infinite_int_iff_unbounded: "infinite S \ (\m. \n. \n\ > m \ n \ S)" for S :: "int set" by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded o_def image_def) (metis (full_types) nat_le_iff nat_mono not_le) proposition finite_int_iff_bounded: "finite S \ (\k. abs ` S \ {.. (\k. abs ` S \ {.. k})" for S :: "int set" using infinite_int_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) subsection \Infinitely Many and Almost All\ text \ We often need to reason about the existence of infinitely many (resp., all but finitely many) objects satisfying some predicate, so we introduce corresponding binders and their proof rules. \ lemma not_INFM [simp]: "\ (INFM x. P x) \ (MOST x. \ P x)" by (rule not_frequently) lemma not_MOST [simp]: "\ (MOST x. P x) \ (INFM x. \ P x)" by (rule not_eventually) lemma INFM_const [simp]: "(INFM x::'a. P) \ P \ infinite (UNIV::'a set)" by (simp add: frequently_const_iff) lemma MOST_const [simp]: "(MOST x::'a. P) \ P \ finite (UNIV::'a set)" by (simp add: eventually_const_iff) lemma INFM_imp_distrib: "(INFM x. P x \ Q x) \ ((MOST x. P x) \ (INFM x. Q x))" by (rule frequently_imp_iff) lemma MOST_imp_iff: "MOST x. P x \ (MOST x. P x \ Q x) \ (MOST x. Q x)" by (auto intro: eventually_rev_mp eventually_mono) lemma INFM_conjI: "INFM x. P x \ MOST x. Q x \ INFM x. P x \ Q x" by (rule frequently_rev_mp[of P]) (auto elim: eventually_mono) text \Properties of quantifiers with injective functions.\ lemma INFM_inj: "INFM x. P (f x) \ inj f \ INFM x. P x" using finite_vimageI[of "{x. P x}" f] by (auto simp: frequently_cofinite) lemma MOST_inj: "MOST x. P x \ inj f \ MOST x. P (f x)" using finite_vimageI[of "{x. \ P x}" f] by (auto simp: eventually_cofinite) text \Properties of quantifiers with singletons.\ lemma not_INFM_eq [simp]: "\ (INFM x. x = a)" "\ (INFM x. a = x)" unfolding frequently_cofinite by simp_all lemma MOST_neq [simp]: "MOST x. x \ a" "MOST x. a \ x" unfolding eventually_cofinite by simp_all lemma INFM_neq [simp]: "(INFM x::'a. x \ a) \ infinite (UNIV::'a set)" "(INFM x::'a. a \ x) \ infinite (UNIV::'a set)" unfolding frequently_cofinite by simp_all lemma MOST_eq [simp]: "(MOST x::'a. x = a) \ finite (UNIV::'a set)" "(MOST x::'a. a = x) \ finite (UNIV::'a set)" unfolding eventually_cofinite by simp_all lemma MOST_eq_imp: "MOST x. x = a \ P x" "MOST x. a = x \ P x" unfolding eventually_cofinite by simp_all text \Properties of quantifiers over the naturals.\ lemma MOST_nat: "(\\<^sub>\n. P n) \ (\m. \n>m. P n)" for P :: "nat \ bool" by (auto simp add: eventually_cofinite finite_nat_iff_bounded_le subset_eq simp flip: not_le) lemma MOST_nat_le: "(\\<^sub>\n. P n) \ (\m. \n\m. P n)" for P :: "nat \ bool" by (auto simp add: eventually_cofinite finite_nat_iff_bounded subset_eq simp flip: not_le) lemma INFM_nat: "(\\<^sub>\n. P n) \ (\m. \n>m. P n)" for P :: "nat \ bool" by (simp add: frequently_cofinite infinite_nat_iff_unbounded) lemma INFM_nat_le: "(\\<^sub>\n. P n) \ (\m. \n\m. P n)" for P :: "nat \ bool" by (simp add: frequently_cofinite infinite_nat_iff_unbounded_le) lemma MOST_INFM: "infinite (UNIV::'a set) \ MOST x::'a. P x \ INFM x::'a. P x" by (simp add: eventually_frequently) lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \ (MOST n. P n)" by (simp add: cofinite_eq_sequentially) lemma MOST_SucI: "MOST n. P n \ MOST n. P (Suc n)" and MOST_SucD: "MOST n. P (Suc n) \ MOST n. P n" by (simp_all add: MOST_Suc_iff) lemma MOST_ge_nat: "MOST n::nat. m \ n" by (simp add: cofinite_eq_sequentially) \ \legacy names\ lemma Inf_many_def: "Inf_many P \ infinite {x. P x}" by (fact frequently_cofinite) lemma Alm_all_def: "Alm_all P \ \ (INFM x. \ P x)" by simp lemma INFM_iff_infinite: "(INFM x. P x) \ infinite {x. P x}" by (fact frequently_cofinite) lemma MOST_iff_cofinite: "(MOST x. P x) \ finite {x. \ P x}" by (fact eventually_cofinite) lemma INFM_EX: "(\\<^sub>\x. P x) \ (\x. P x)" by (fact frequently_ex) lemma ALL_MOST: "\x. P x \ \\<^sub>\x. P x" by (fact always_eventually) lemma INFM_mono: "\\<^sub>\x. P x \ (\x. P x \ Q x) \ \\<^sub>\x. Q x" by (fact frequently_elim1) lemma MOST_mono: "\\<^sub>\x. P x \ (\x. P x \ Q x) \ \\<^sub>\x. Q x" by (fact eventually_mono) lemma INFM_disj_distrib: "(\\<^sub>\x. P x \ Q x) \ (\\<^sub>\x. P x) \ (\\<^sub>\x. Q x)" by (fact frequently_disj_iff) lemma MOST_rev_mp: "\\<^sub>\x. P x \ \\<^sub>\x. P x \ Q x \ \\<^sub>\x. Q x" by (fact eventually_rev_mp) lemma MOST_conj_distrib: "(\\<^sub>\x. P x \ Q x) \ (\\<^sub>\x. P x) \ (\\<^sub>\x. Q x)" by (fact eventually_conj_iff) lemma MOST_conjI: "MOST x. P x \ MOST x. Q x \ MOST x. P x \ Q x" by (fact eventually_conj) lemma INFM_finite_Bex_distrib: "finite A \ (INFM y. \x\A. P x y) \ (\x\A. INFM y. P x y)" by (fact frequently_bex_finite_distrib) lemma MOST_finite_Ball_distrib: "finite A \ (MOST y. \x\A. P x y) \ (\x\A. MOST y. P x y)" by (fact eventually_ball_finite_distrib) lemma INFM_E: "INFM x. P x \ (\x. P x \ thesis) \ thesis" by (fact frequentlyE) lemma MOST_I: "(\x. P x) \ MOST x. P x" by (rule eventuallyI) lemmas MOST_iff_finiteNeg = MOST_iff_cofinite subsection \Enumeration of an Infinite Set\ text \The set's element type must be wellordered (e.g. the natural numbers).\ text \ Could be generalized to \<^prop>\enumerate' S n = (SOME t. t \ s \ finite {s\S. s < t} \ card {s\S. s < t} = n)\. \ primrec (in wellorder) enumerate :: "'a set \ nat \ 'a" where enumerate_0: "enumerate S 0 = (LEAST n. n \ S)" | enumerate_Suc: "enumerate S (Suc n) = enumerate (S - {LEAST n. n \ S}) n" lemma enumerate_Suc': "enumerate S (Suc n) = enumerate (S - {enumerate S 0}) n" by simp lemma enumerate_in_set: "infinite S \ enumerate S n \ S" proof (induct n arbitrary: S) case 0 then show ?case by (fastforce intro: LeastI dest!: infinite_imp_nonempty) next case (Suc n) then show ?case by simp (metis DiffE infinite_remove) qed declare enumerate_0 [simp del] enumerate_Suc [simp del] lemma enumerate_step: "infinite S \ enumerate S n < enumerate S (Suc n)" proof (induction n arbitrary: S) case 0 then have "enumerate S 0 \ enumerate S (Suc 0)" by (simp add: enumerate_0 Least_le enumerate_in_set) moreover have "enumerate (S - {enumerate S 0}) 0 \ S - {enumerate S 0}" by (meson "0.prems" enumerate_in_set infinite_remove) then have "enumerate S 0 \ enumerate (S - {enumerate S 0}) 0" by auto ultimately show ?case by (simp add: enumerate_Suc') next case (Suc n) then show ?case by (simp add: enumerate_Suc') qed lemma enumerate_mono: "m < n \ infinite S \ enumerate S m < enumerate S n" by (induct m n rule: less_Suc_induct) (auto intro: enumerate_step) lemma enumerate_mono_iff [simp]: "infinite S \ enumerate S m < enumerate S n \ m < n" by (metis enumerate_mono less_asym less_linear) lemma le_enumerate: assumes S: "infinite S" shows "n \ enumerate S n" using S proof (induct n) case 0 then show ?case by simp next case (Suc n) then have "n \ enumerate S n" by simp also note enumerate_mono[of n "Suc n", OF _ \infinite S\] finally show ?case by simp qed lemma infinite_enumerate: assumes fS: "infinite S" shows "\r::nat\nat. strict_mono r \ (\n. r n \ S)" unfolding strict_mono_def using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by blast lemma enumerate_Suc'': fixes S :: "'a::wellorder set" assumes "infinite S" shows "enumerate S (Suc n) = (LEAST s. s \ S \ enumerate S n < s)" using assms proof (induct n arbitrary: S) case 0 then have "\s \ S. enumerate S 0 \ s" by (auto simp: enumerate.simps intro: Least_le) then show ?case unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"] by (intro arg_cong[where f = Least] ext) auto next case (Suc n S) show ?case using enumerate_mono[OF zero_less_Suc \infinite S\, of n] \infinite S\ apply (subst (1 2) enumerate_Suc') apply (subst Suc) apply (use \infinite S\ in simp) apply (intro arg_cong[where f = Least] ext) apply (auto simp flip: enumerate_Suc') done qed lemma enumerate_Ex: fixes S :: "nat set" assumes S: "infinite S" and s: "s \ S" shows "\n. enumerate S n = s" using s proof (induct s rule: less_induct) case (less s) show ?case proof (cases "\y\S. y < s") case True let ?y = "Max {s'\S. s' < s}" from True have y: "\x. ?y < x \ (\s'\S. s' < s \ s' < x)" by (subst Max_less_iff) auto then have y_in: "?y \ {s'\S. s' < s}" by (intro Max_in) auto with less.hyps[of ?y] obtain n where "enumerate S n = ?y" by auto with S have "enumerate S (Suc n) = s" by (auto simp: y less enumerate_Suc'' intro!: Least_equality) then show ?thesis by auto next case False then have "\t\S. s \ t" by auto with \s \ S\ show ?thesis by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) qed qed lemma inj_enumerate: fixes S :: "'a::wellorder set" assumes S: "infinite S" shows "inj (enumerate S)" unfolding inj_on_def proof clarsimp show "\x y. enumerate S x = enumerate S y \ x = y" by (metis neq_iff enumerate_mono[OF _ \infinite S\]) qed text \To generalise this, we'd need a condition that all initial segments were finite\ lemma bij_enumerate: fixes S :: "nat set" assumes S: "infinite S" shows "bij_betw (enumerate S) UNIV S" proof - have "\s \ S. \i. enumerate S i = s" using enumerate_Ex[OF S] by auto moreover note \infinite S\ inj_enumerate ultimately show ?thesis unfolding bij_betw_def by (auto intro: enumerate_in_set) qed +lemma + fixes S :: "nat set" + assumes S: "infinite S" + shows range_enumerate: "range (enumerate S) = S" + and strict_mono_enumerate: "strict_mono (enumerate S)" + by (auto simp add: bij_betw_imp_surj_on bij_enumerate assms strict_mono_def) + text \A pair of weird and wonderful lemmas from HOL Light.\ lemma finite_transitivity_chain: assumes "finite A" and R: "\x. \ R x x" "\x y z. \R x y; R y z\ \ R x z" and A: "\x. x \ A \ \y. y \ A \ R x y" shows "A = {}" using \finite A\ A proof (induct A) case empty then show ?case by simp next case (insert a A) have False using R(1)[of a] R(2)[of _ a] insert(3,4) by blast thus ?case .. qed corollary Union_maximal_sets: assumes "finite \" shows "\{T \ \. \U\\. \ T \ U} = \\" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by force show "?rhs \ ?lhs" proof (rule Union_subsetI) fix S assume "S \ \" have "{T \ \. S \ T} = {}" if "\ (\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y)" proof - have \
: "\x. x \ \ \ S \ x \ \y. y \ \ \ S \ y \ x \ y" using that by (blast intro: dual_order.trans psubset_imp_subset) show ?thesis proof (rule finite_transitivity_chain [of _ "\T U. S \ T \ T \ U"]) qed (use assms in \auto intro: \
\) qed with \S \ \\ show "\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y" by blast qed qed subsection \Properties of @{term enumerate} on finite sets\ lemma finite_enumerate_in_set: "\finite S; n < card S\ \ enumerate S n \ S" proof (induction n arbitrary: S) case 0 then show ?case by (metis all_not_in_conv card_empty enumerate.simps(1) not_less0 wellorder_Least_lemma(1)) next case (Suc n) show ?case using Suc.prems Suc.IH [of "S - {LEAST n. n \ S}"] apply (simp add: enumerate.simps) by (metis Diff_empty Diff_insert0 Suc_lessD card.remove less_Suc_eq) qed lemma finite_enumerate_step: "\finite S; Suc n < card S\ \ enumerate S n < enumerate S (Suc n)" proof (induction n arbitrary: S) case 0 then have "enumerate S 0 \ enumerate S (Suc 0)" by (simp add: Least_le enumerate.simps(1) finite_enumerate_in_set) moreover have "enumerate (S - {enumerate S 0}) 0 \ S - {enumerate S 0}" by (metis 0 Suc_lessD Suc_less_eq card_Suc_Diff1 enumerate_in_set finite_enumerate_in_set) then have "enumerate S 0 \ enumerate (S - {enumerate S 0}) 0" by auto ultimately show ?case by (simp add: enumerate_Suc') next case (Suc n) then show ?case by (simp add: enumerate_Suc' finite_enumerate_in_set) qed lemma finite_enumerate_mono: "\m < n; finite S; n < card S\ \ enumerate S m < enumerate S n" by (induct m n rule: less_Suc_induct) (auto intro: finite_enumerate_step) lemma finite_enumerate_mono_iff [simp]: "\finite S; m < card S; n < card S\ \ enumerate S m < enumerate S n \ m < n" by (metis finite_enumerate_mono less_asym less_linear) lemma finite_le_enumerate: assumes "finite S" "n < card S" shows "n \ enumerate S n" using assms proof (induction n) case 0 then show ?case by simp next case (Suc n) then have "n \ enumerate S n" by simp also note finite_enumerate_mono[of n "Suc n", OF _ \finite S\] finally show ?case using Suc.prems(2) Suc_leI by blast qed lemma finite_enumerate: assumes fS: "finite S" shows "\r::nat\nat. strict_mono_on r {.. (\n S)" unfolding strict_mono_def using finite_enumerate_in_set[OF fS] finite_enumerate_mono[of _ _ S] fS by (metis lessThan_iff strict_mono_on_def) lemma finite_enumerate_Suc'': fixes S :: "'a::wellorder set" assumes "finite S" "Suc n < card S" shows "enumerate S (Suc n) = (LEAST s. s \ S \ enumerate S n < s)" using assms proof (induction n arbitrary: S) case 0 then have "\s \ S. enumerate S 0 \ s" by (auto simp: enumerate.simps intro: Least_le) then show ?case unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"] by (metis Diff_iff dual_order.strict_iff_order singletonD singletonI) next case (Suc n S) then have "Suc n < card (S - {enumerate S 0})" using Suc.prems(2) finite_enumerate_in_set by force then show ?case apply (subst (1 2) enumerate_Suc') apply (simp add: Suc) apply (intro arg_cong[where f = Least] HOL.ext) using finite_enumerate_mono[OF zero_less_Suc \finite S\, of n] Suc.prems by (auto simp flip: enumerate_Suc') qed lemma finite_enumerate_initial_segment: fixes S :: "'a::wellorder set" assumes "finite S" and n: "n < card (S \ {.. {.. S \ n < s) = (LEAST n. n \ S)" proof (rule Least_equality) have "\t. t \ S \ t < s" by (metis "0" card_gt_0_iff disjoint_iff_not_equal lessThan_iff) then show "(LEAST n. n \ S) \ S \ (LEAST n. n \ S) < s" by (meson LeastI Least_le le_less_trans) qed (simp add: Least_le) then show ?case by (auto simp: enumerate_0) next case (Suc n) then have less_card: "Suc n < card S" by (meson assms(1) card_mono inf_sup_ord(1) leD le_less_linear order.trans) obtain T where T: "T \ {s \ S. enumerate S n < s}" by (metis Infinite_Set.enumerate_step enumerate_in_set finite_enumerate_in_set finite_enumerate_step less_card mem_Collect_eq) have "(LEAST x. x \ S \ x < s \ enumerate S n < x) = (LEAST x. x \ S \ enumerate S n < x)" (is "_ = ?r") proof (intro Least_equality conjI) show "?r \ S" by (metis (mono_tags, lifting) LeastI mem_Collect_eq T) have "\ s \ ?r" using not_less_Least [of _ "\x. x \ S \ enumerate S n < x"] Suc assms by (metis (mono_tags, lifting) Int_Collect Suc_lessD finite_Int finite_enumerate_in_set finite_enumerate_step lessThan_def less_le_trans) then show "?r < s" by auto show "enumerate S n < ?r" by (metis (no_types, lifting) LeastI mem_Collect_eq T) qed (auto simp: Least_le) then show ?case using Suc assms by (simp add: finite_enumerate_Suc'' less_card) qed lemma finite_enumerate_Ex: fixes S :: "'a::wellorder set" assumes S: "finite S" and s: "s \ S" shows "\ny\S. y < s") case True let ?T = "S \ {..finite S\]) from True have y: "\x. Max ?T < x \ (\s'\S. s' < s \ s' < x)" by (subst Max_less_iff) (auto simp: \finite ?T\) then have y_in: "Max ?T \ {s'\S. s' < s}" using Max_in \finite ?T\ by fastforce with less.IH[of "Max ?T" ?T] obtain n where n: "enumerate ?T n = Max ?T" "n < card ?T" using \finite ?T\ by blast then have "Suc n < card S" using TS less_trans_Suc by blast with S n have "enumerate S (Suc n) = s" by (subst finite_enumerate_Suc'') (auto simp: y finite_enumerate_initial_segment less finite_enumerate_Suc'' intro!: Least_equality) then show ?thesis using \Suc n < card S\ by blast next case False then have "\t\S. s \ t" by auto moreover have "0 < card S" using card_0_eq less.prems by blast ultimately show ?thesis using \s \ S\ by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) qed qed lemma finite_enum_subset: assumes "\i. i < card X \ enumerate X i = enumerate Y i" and "finite X" "finite Y" "card X \ card Y" shows "X \ Y" by (metis assms finite_enumerate_Ex finite_enumerate_in_set less_le_trans subsetI) lemma finite_enum_ext: assumes "\i. i < card X \ enumerate X i = enumerate Y i" and "finite X" "finite Y" "card X = card Y" shows "X = Y" by (intro antisym finite_enum_subset) (auto simp: assms) lemma finite_bij_enumerate: fixes S :: "'a::wellorder set" assumes S: "finite S" shows "bij_betw (enumerate S) {..n m. \n \ m; n < card S; m < card S\ \ enumerate S n \ enumerate S m" using finite_enumerate_mono[OF _ \finite S\] by (auto simp: neq_iff) then have "inj_on (enumerate S) {..s \ S. \ifinite S\ ultimately show ?thesis unfolding bij_betw_def by (auto intro: finite_enumerate_in_set) qed lemma ex_bij_betw_strict_mono_card: fixes M :: "'a::wellorder set" assumes "finite M" obtains h where "bij_betw h {..