clasohm@1465: (* Title: HOL/Finite.thy clasohm@923: ID: $Id$ nipkow@1531: Author: Lawrence C Paulson & Tobias Nipkow nipkow@1531: Copyright 1995 University of Cambridge & TU Muenchen clasohm@923: nipkow@1531: Finite sets and their cardinality clasohm@923: *) clasohm@923: clasohm@923: open Finite; clasohm@923: nipkow@3413: section "finite"; nipkow@1531: nipkow@3413: (* clasohm@923: goalw Finite.thy Fin.defs "!!A B. A<=B ==> Fin(A) <= Fin(B)"; clasohm@1465: by (rtac lfp_mono 1); clasohm@923: by (REPEAT (ares_tac basic_monos 1)); clasohm@923: qed "Fin_mono"; clasohm@923: clasohm@923: goalw Finite.thy Fin.defs "Fin(A) <= Pow(A)"; wenzelm@4089: by (blast_tac (claset() addSIs [lfp_lowerbound]) 1); clasohm@923: qed "Fin_subset_Pow"; clasohm@923: clasohm@923: (* A : Fin(B) ==> A <= B *) clasohm@923: val FinD = Fin_subset_Pow RS subsetD RS PowD; nipkow@3413: *) clasohm@923: clasohm@923: (*Discharging ~ x:y entails extra work*) clasohm@923: val major::prems = goal Finite.thy nipkow@3413: "[| finite F; P({}); \ nipkow@3413: \ !!F x. [| finite F; x ~: F; P(F) |] ==> P(insert x F) \ clasohm@923: \ |] ==> P(F)"; nipkow@3413: by (rtac (major RS Finites.induct) 1); nipkow@3413: by (excluded_middle_tac "a:A" 2); clasohm@923: by (etac (insert_absorb RS ssubst) 3 THEN assume_tac 3); (*backtracking!*) clasohm@923: by (REPEAT (ares_tac prems 1)); nipkow@3413: qed "finite_induct"; nipkow@3413: nipkow@3413: val major::prems = goal Finite.thy nipkow@3413: "[| finite F; \ nipkow@3413: \ P({}); \ nipkow@3413: \ !!F a. [| finite F; a:A; a ~: F; P(F) |] ==> P(insert a F) \ nipkow@3413: \ |] ==> F <= A --> P(F)"; nipkow@3413: by (rtac (major RS finite_induct) 1); wenzelm@4089: by (ALLGOALS (blast_tac (claset() addIs prems))); nipkow@3413: val lemma = result(); clasohm@923: nipkow@3413: val prems = goal Finite.thy nipkow@3413: "[| finite F; F <= A; \ nipkow@3413: \ P({}); \ nipkow@3413: \ !!F a. [| finite F; a:A; a ~: F; P(F) |] ==> P(insert a F) \ nipkow@3413: \ |] ==> P(F)"; paulson@3457: by (blast_tac (HOL_cs addIs ((lemma RS mp)::prems)) 1); nipkow@3413: qed "finite_subset_induct"; nipkow@3413: nipkow@3413: Addsimps Finites.intrs; nipkow@3413: AddSIs Finites.intrs; clasohm@923: clasohm@923: (*The union of two finite sets is finite*) clasohm@923: val major::prems = goal Finite.thy nipkow@3413: "[| finite F; finite G |] ==> finite(F Un G)"; nipkow@3413: by (rtac (major RS finite_induct) 1); wenzelm@4089: by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); nipkow@3413: qed "finite_UnI"; clasohm@923: clasohm@923: (*Every subset of a finite set is finite*) nipkow@3413: val [subs,fin] = goal Finite.thy "[| A<=B; finite B |] ==> finite A"; nipkow@3413: by (EVERY1 [subgoal_tac "ALL C. C<=B --> finite C", clasohm@1465: rtac mp, etac spec, clasohm@1465: rtac subs]); nipkow@3413: by (rtac (fin RS finite_induct) 1); wenzelm@4089: by (simp_tac (simpset() addsimps [subset_Un_eq]) 1); wenzelm@4089: by (safe_tac (claset() addSDs [subset_insert_iff RS iffD1])); clasohm@923: by (eres_inst_tac [("t","C")] (insert_Diff RS subst) 2); clasohm@1264: by (ALLGOALS Asm_simp_tac); nipkow@3413: qed "finite_subset"; clasohm@923: nipkow@3413: goal Finite.thy "finite(F Un G) = (finite F & finite G)"; wenzelm@4089: by (blast_tac (claset() addIs [finite_UnI] addDs nipkow@3413: [Un_upper1 RS finite_subset, Un_upper2 RS finite_subset]) 1); nipkow@3413: qed "finite_Un"; nipkow@3413: AddIffs[finite_Un]; nipkow@1531: nipkow@3413: goal Finite.thy "finite(insert a A) = finite A"; paulson@1553: by (stac insert_is_Un 1); nipkow@3413: by (simp_tac (HOL_ss addsimps [finite_Un]) 1); paulson@3427: by (Blast_tac 1); nipkow@3413: qed "finite_insert"; nipkow@3413: Addsimps[finite_insert]; nipkow@1531: nipkow@3413: (*The image of a finite set is finite *) nipkow@3413: goal Finite.thy "!!F. finite F ==> finite(h``F)"; nipkow@3413: by (etac finite_induct 1); clasohm@1264: by (Simp_tac 1); nipkow@3413: by (Asm_simp_tac 1); nipkow@3413: qed "finite_imageI"; clasohm@923: clasohm@923: val major::prems = goal Finite.thy nipkow@3413: "[| finite c; finite b; \ clasohm@1465: \ P(b); \ nipkow@3413: \ !!x y. [| finite y; x:y; P(y) |] ==> P(y-{x}) \ clasohm@923: \ |] ==> c<=b --> P(b-c)"; nipkow@3413: by (rtac (major RS finite_induct) 1); paulson@2031: by (stac Diff_insert 2); clasohm@923: by (ALLGOALS (asm_simp_tac wenzelm@4089: (simpset() addsimps (prems@[Diff_subset RS finite_subset])))); nipkow@1531: val lemma = result(); clasohm@923: clasohm@923: val prems = goal Finite.thy nipkow@3413: "[| finite A; \ nipkow@3413: \ P(A); \ nipkow@3413: \ !!a A. [| finite A; a:A; P(A) |] ==> P(A-{a}) \ clasohm@923: \ |] ==> P({})"; clasohm@923: by (rtac (Diff_cancel RS subst) 1); nipkow@1531: by (rtac (lemma RS mp) 1); clasohm@923: by (REPEAT (ares_tac (subset_refl::prems) 1)); nipkow@3413: qed "finite_empty_induct"; nipkow@1531: nipkow@1531: paulson@1618: (* finite B ==> finite (B - Ba) *) paulson@1618: bind_thm ("finite_Diff", Diff_subset RS finite_subset); nipkow@1531: Addsimps [finite_Diff]; nipkow@1531: paulson@3368: goal Finite.thy "finite(A-{a}) = finite(A)"; paulson@3368: by (case_tac "a:A" 1); paulson@3457: by (rtac (finite_insert RS sym RS trans) 1); paulson@3368: by (stac insert_Diff 1); paulson@3368: by (ALLGOALS Asm_simp_tac); paulson@3368: qed "finite_Diff_singleton"; paulson@3368: AddIffs [finite_Diff_singleton]; paulson@3368: paulson@4059: (*Lemma for proving finite_imageD*) nipkow@3413: goal Finite.thy "!!A. finite B ==> !A. f``A = B --> inj_onto f A --> finite A"; paulson@1553: by (etac finite_induct 1); nipkow@3413: by (ALLGOALS Asm_simp_tac); paulson@3708: by (Clarify_tac 1); nipkow@3413: by (subgoal_tac "EX y:A. f y = x & F = f``(A-{y})" 1); paulson@3708: by (Clarify_tac 1); wenzelm@4089: by (full_simp_tac (simpset() addsimps [inj_onto_def]) 1); nipkow@3413: by (Blast_tac 1); paulson@3368: by (thin_tac "ALL A. ?PP(A)" 1); nipkow@3413: by (forward_tac [[equalityD2, insertI1] MRS subsetD] 1); paulson@3708: by (Clarify_tac 1); paulson@3368: by (res_inst_tac [("x","xa")] bexI 1); paulson@4059: by (ALLGOALS wenzelm@4089: (asm_full_simp_tac (simpset() addsimps [inj_onto_image_set_diff]))); paulson@3368: val lemma = result(); paulson@3368: paulson@3368: goal Finite.thy "!!A. [| finite(f``A); inj_onto f A |] ==> finite A"; paulson@3457: by (dtac lemma 1); paulson@3368: by (Blast_tac 1); paulson@3368: qed "finite_imageD"; paulson@3368: nipkow@4014: (** The finite UNION of finite sets **) nipkow@4014: nipkow@4014: val [prem] = goal Finite.thy nipkow@4014: "finite A ==> (!a:A. finite(B a)) --> finite(UN a:A. B a)"; paulson@4153: by (rtac (prem RS finite_induct) 1); paulson@4153: by (ALLGOALS Asm_simp_tac); nipkow@4014: bind_thm("finite_UnionI", ballI RSN (2, result() RS mp)); nipkow@4014: Addsimps [finite_UnionI]; nipkow@4014: nipkow@4014: (** Sigma of finite sets **) nipkow@4014: nipkow@4014: goalw Finite.thy [Sigma_def] nipkow@4014: "!!A. [| finite A; !a:A. finite(B a) |] ==> finite(SIGMA a:A. B a)"; paulson@4153: by (blast_tac (claset() addSIs [finite_UnionI]) 1); nipkow@4014: bind_thm("finite_SigmaI", ballI RSN (2,result())); nipkow@4014: Addsimps [finite_SigmaI]; paulson@3368: paulson@3368: (** The powerset of a finite set **) paulson@3368: paulson@3368: goal Finite.thy "!!A. finite(Pow A) ==> finite A"; paulson@3368: by (subgoal_tac "finite ((%x.{x})``A)" 1); paulson@3457: by (rtac finite_subset 2); paulson@3457: by (assume_tac 3); paulson@3368: by (ALLGOALS wenzelm@4089: (fast_tac (claset() addSDs [rewrite_rule [inj_onto_def] finite_imageD]))); paulson@3368: val lemma = result(); paulson@3368: paulson@3368: goal Finite.thy "finite(Pow A) = finite A"; paulson@3457: by (rtac iffI 1); paulson@3457: by (etac lemma 1); paulson@3368: (*Opposite inclusion: finite A ==> finite (Pow A) *) paulson@3340: by (etac finite_induct 1); paulson@3340: by (ALLGOALS paulson@3340: (asm_simp_tac wenzelm@4089: (simpset() addsimps [finite_UnI, finite_imageI, Pow_insert]))); paulson@3368: qed "finite_Pow_iff"; paulson@3368: AddIffs [finite_Pow_iff]; paulson@3340: nipkow@3439: goal Finite.thy "finite(r^-1) = finite r"; paulson@3457: by (subgoal_tac "r^-1 = (%(x,y).(y,x))``r" 1); paulson@3457: by (Asm_simp_tac 1); paulson@3457: by (rtac iffI 1); paulson@3457: by (etac (rewrite_rule [inj_onto_def] finite_imageD) 1); wenzelm@4089: by (simp_tac (simpset() addsplits [expand_split]) 1); paulson@3457: by (etac finite_imageI 1); wenzelm@4089: by (simp_tac (simpset() addsimps [inverse_def,image_def]) 1); paulson@3457: by (Auto_tac()); paulson@3457: by (rtac bexI 1); paulson@3457: by (assume_tac 2); paulson@3457: by (Simp_tac 1); paulson@3457: by (split_all_tac 1); paulson@3457: by (Asm_full_simp_tac 1); nipkow@3439: qed "finite_inverse"; nipkow@3439: AddIffs [finite_inverse]; nipkow@1531: nipkow@1548: section "Finite cardinality -- 'card'"; nipkow@1531: nipkow@1531: goal Set.thy "{f i |i. P i | i=n} = insert (f n) {f i|i. P i}"; paulson@2922: by (Blast_tac 1); nipkow@1531: val Collect_conv_insert = result(); nipkow@1531: nipkow@1531: goalw Finite.thy [card_def] "card {} = 0"; paulson@1553: by (rtac Least_equality 1); paulson@1553: by (ALLGOALS Asm_full_simp_tac); nipkow@1531: qed "card_empty"; nipkow@1531: Addsimps [card_empty]; nipkow@1531: nipkow@1531: val [major] = goal Finite.thy nipkow@1531: "finite A ==> ? (n::nat) f. A = {f i |i. i \ wenzelm@3842: \ ? m::nat. m \ wenzelm@3842: \ (LEAST n. ? f. insert x A = {f i|i. i card(insert x A) = Suc(card A)"; paulson@1553: by (etac lemma 1); paulson@1553: by (assume_tac 1); nipkow@1531: qed "card_insert_disjoint"; paulson@3352: Addsimps [card_insert_disjoint]; paulson@3352: paulson@3352: goal Finite.thy "!!A. finite A ==> !B. B <= A --> card(B) <= card(A)"; paulson@3352: by (etac finite_induct 1); paulson@3352: by (Simp_tac 1); paulson@3708: by (Clarify_tac 1); paulson@3352: by (case_tac "x:B" 1); nipkow@3413: by (dres_inst_tac [("A","B")] mk_disjoint_insert 1); paulson@4153: by (SELECT_GOAL Safe_tac 1); paulson@3352: by (rotate_tac ~1 1); wenzelm@4089: by (asm_full_simp_tac (simpset() addsimps [subset_insert_iff,finite_subset]) 1); paulson@3352: by (rotate_tac ~1 1); wenzelm@4089: by (asm_full_simp_tac (simpset() addsimps [subset_insert_iff,finite_subset]) 1); paulson@3352: qed_spec_mp "card_mono"; paulson@3352: paulson@3352: goal Finite.thy "!!A B. [| finite A; finite B |]\ paulson@3352: \ ==> A Int B = {} --> card(A Un B) = card A + card B"; paulson@3352: by (etac finite_induct 1); paulson@3352: by (ALLGOALS wenzelm@4089: (asm_simp_tac (simpset() addsimps [Int_insert_left] nipkow@3919: addsplits [expand_if]))); paulson@3352: qed_spec_mp "card_Un_disjoint"; paulson@3352: paulson@3352: goal Finite.thy "!!A. [| finite A; B<=A |] ==> card A - card B = card (A - B)"; paulson@3352: by (subgoal_tac "(A-B) Un B = A" 1); paulson@3352: by (Blast_tac 2); paulson@3457: by (rtac (add_right_cancel RS iffD1) 1); paulson@3457: by (rtac (card_Un_disjoint RS subst) 1); paulson@3457: by (etac ssubst 4); paulson@3352: by (Blast_tac 3); paulson@3352: by (ALLGOALS paulson@3352: (asm_simp_tac wenzelm@4089: (simpset() addsimps [add_commute, not_less_iff_le, paulson@3352: add_diff_inverse, card_mono, finite_subset]))); paulson@3352: qed "card_Diff_subset"; nipkow@1531: paulson@1618: goal Finite.thy "!!A. [| finite A; x: A |] ==> Suc(card(A-{x})) = card A"; paulson@1618: by (res_inst_tac [("t", "A")] (insert_Diff RS subst) 1); paulson@1618: by (assume_tac 1); paulson@3352: by (Asm_simp_tac 1); paulson@1618: qed "card_Suc_Diff"; paulson@1618: paulson@1618: goal Finite.thy "!!A. [| finite A; x: A |] ==> card(A-{x}) < card A"; paulson@2031: by (rtac Suc_less_SucD 1); wenzelm@4089: by (asm_simp_tac (simpset() addsimps [card_Suc_Diff]) 1); paulson@1618: qed "card_Diff"; paulson@1618: paulson@3389: paulson@3389: (*** Cardinality of the Powerset ***) paulson@3389: nipkow@1531: val [major] = goal Finite.thy nipkow@1531: "finite A ==> card(insert x A) = Suc(card(A-{x}))"; paulson@1553: by (case_tac "x:A" 1); wenzelm@4089: by (asm_simp_tac (simpset() addsimps [insert_absorb]) 1); paulson@1553: by (dtac mk_disjoint_insert 1); paulson@1553: by (etac exE 1); paulson@1553: by (Asm_simp_tac 1); paulson@1553: by (rtac card_insert_disjoint 1); paulson@1553: by (rtac (major RSN (2,finite_subset)) 1); paulson@2922: by (Blast_tac 1); paulson@2922: by (Blast_tac 1); wenzelm@4089: by (asm_simp_tac (simpset() addsimps [major RS card_insert_disjoint]) 1); nipkow@1531: qed "card_insert"; nipkow@1531: Addsimps [card_insert]; nipkow@1531: paulson@3340: goal Finite.thy "!!A. finite(A) ==> inj_onto f A --> card (f `` A) = card A"; paulson@3340: by (etac finite_induct 1); paulson@3340: by (ALLGOALS Asm_simp_tac); paulson@3724: by Safe_tac; paulson@3457: by (rewtac inj_onto_def); paulson@3340: by (Blast_tac 1); paulson@3340: by (stac card_insert_disjoint 1); paulson@3340: by (etac finite_imageI 1); paulson@3340: by (Blast_tac 1); paulson@3340: by (Blast_tac 1); paulson@3340: qed_spec_mp "card_image"; paulson@3340: paulson@3389: goal thy "!!A. finite A ==> card (Pow A) = 2 ^ card A"; paulson@3389: by (etac finite_induct 1); wenzelm@4089: by (ALLGOALS (asm_simp_tac (simpset() addsimps [Pow_insert]))); paulson@3389: by (stac card_Un_disjoint 1); wenzelm@4089: by (EVERY (map (blast_tac (claset() addIs [finite_imageI])) [3,2,1])); paulson@3389: by (subgoal_tac "inj_onto (insert x) (Pow F)" 1); wenzelm@4089: by (asm_simp_tac (simpset() addsimps [card_image, Pow_insert]) 1); paulson@3457: by (rewtac inj_onto_def); wenzelm@4089: by (blast_tac (claset() addSEs [equalityE]) 1); paulson@3389: qed "card_Pow"; paulson@3389: Addsimps [card_Pow]; paulson@3340: paulson@3389: paulson@3389: (*Proper subsets*) nipkow@3222: goalw Finite.thy [psubset_def] nipkow@3222: "!!B. finite B ==> !A. A < B --> card(A) < card(B)"; nipkow@3222: by (etac finite_induct 1); nipkow@3222: by (Simp_tac 1); paulson@3708: by (Clarify_tac 1); nipkow@3222: by (case_tac "x:A" 1); nipkow@3222: (*1*) nipkow@3413: by (dres_inst_tac [("A","A")]mk_disjoint_insert 1); nipkow@3222: by (etac exE 1); nipkow@3222: by (etac conjE 1); nipkow@3222: by (hyp_subst_tac 1); nipkow@3222: by (rotate_tac ~1 1); wenzelm@4089: by (asm_full_simp_tac (simpset() addsimps [subset_insert_iff,finite_subset]) 1); paulson@3708: by (Blast_tac 1); nipkow@3222: (*2*) nipkow@3222: by (rotate_tac ~1 1); paulson@3708: by (eres_inst_tac [("P","?a finite (Union C) --> \ paulson@3368: \ (! c : C. k dvd card c) --> \ paulson@3368: \ (! c1: C. ! c2: C. c1 ~= c2 --> c1 Int c2 = {}) \ paulson@3368: \ --> k dvd card(Union C)"; paulson@3368: by (etac finite_induct 1); paulson@3368: by (ALLGOALS Asm_simp_tac); paulson@3708: by (Clarify_tac 1); paulson@3368: by (stac card_Un_disjoint 1); paulson@3368: by (ALLGOALS wenzelm@4089: (asm_full_simp_tac (simpset() paulson@3368: addsimps [dvd_add, disjoint_eq_subset_Compl]))); paulson@3368: by (thin_tac "!c:F. ?PP(c)" 1); paulson@3368: by (thin_tac "!c:F. ?PP(c) & ?QQ(c)" 1); paulson@3708: by (Clarify_tac 1); paulson@3368: by (ball_tac 1); paulson@3368: by (Blast_tac 1); paulson@3368: qed_spec_mp "dvd_partition"; paulson@3368: