src/HOL/Finite_Set.thy
changeset 12396 2298d5b8e530
child 12693 827818b891c7
equal deleted inserted replaced
12395:d6913de7655f 12396:2298d5b8e530
       
     1 (*  Title:      HOL/Finite_Set.thy
       
     2     ID:         $Id$
       
     3     Author:     Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
       
     4     License:    GPL (GNU GENERAL PUBLIC LICENSE)
       
     5 *)
       
     6 
       
     7 header {* Finite sets *}
       
     8 
       
     9 theory Finite_Set = Divides + Power + Inductive + SetInterval:
       
    10 
       
    11 subsection {* Collection of finite sets *}
       
    12 
       
    13 consts Finites :: "'a set set"
       
    14 
       
    15 inductive Finites
       
    16   intros
       
    17     emptyI [simp, intro!]: "{} : Finites"
       
    18     insertI [simp, intro!]: "A : Finites ==> insert a A : Finites"
       
    19 
       
    20 syntax
       
    21   finite :: "'a set => bool"
       
    22 translations
       
    23   "finite A" == "A : Finites"
       
    24 
       
    25 axclass finite \<subseteq> type
       
    26   finite: "finite UNIV"
       
    27 
       
    28 
       
    29 lemma finite_induct [case_names empty insert, induct set: Finites]:
       
    30   "finite F ==>
       
    31     P {} ==> (!!F x. finite F ==> x \<notin> F ==> P F ==> P (insert x F)) ==> P F"
       
    32   -- {* Discharging @{text "x \<notin> F"} entails extra work. *}
       
    33 proof -
       
    34   assume "P {}" and insert: "!!F x. finite F ==> x \<notin> F ==> P F ==> P (insert x F)"
       
    35   assume "finite F"
       
    36   thus "P F"
       
    37   proof induct
       
    38     show "P {}" .
       
    39     fix F x assume F: "finite F" and P: "P F"
       
    40     show "P (insert x F)"
       
    41     proof cases
       
    42       assume "x \<in> F"
       
    43       hence "insert x F = F" by (rule insert_absorb)
       
    44       with P show ?thesis by (simp only:)
       
    45     next
       
    46       assume "x \<notin> F"
       
    47       from F this P show ?thesis by (rule insert)
       
    48     qed
       
    49   qed
       
    50 qed
       
    51 
       
    52 lemma finite_subset_induct [consumes 2, case_names empty insert]:
       
    53   "finite F ==> F \<subseteq> A ==>
       
    54     P {} ==> (!!F a. finite F ==> a \<in> A ==> a \<notin> F ==> P F ==> P (insert a F)) ==>
       
    55     P F"
       
    56 proof -
       
    57   assume "P {}" and insert: "!!F a. finite F ==> a \<in> A ==> a \<notin> F ==> P F ==> P (insert a F)"
       
    58   assume "finite F"
       
    59   thus "F \<subseteq> A ==> P F"
       
    60   proof induct
       
    61     show "P {}" .
       
    62     fix F x assume "finite F" and "x \<notin> F"
       
    63       and P: "F \<subseteq> A ==> P F" and i: "insert x F \<subseteq> A"
       
    64     show "P (insert x F)"
       
    65     proof (rule insert)
       
    66       from i show "x \<in> A" by blast
       
    67       from i have "F \<subseteq> A" by blast
       
    68       with P show "P F" .
       
    69     qed
       
    70   qed
       
    71 qed
       
    72 
       
    73 lemma finite_UnI: "finite F ==> finite G ==> finite (F Un G)"
       
    74   -- {* The union of two finite sets is finite. *}
       
    75   by (induct set: Finites) simp_all
       
    76 
       
    77 lemma finite_subset: "A \<subseteq> B ==> finite B ==> finite A"
       
    78   -- {* Every subset of a finite set is finite. *}
       
    79 proof -
       
    80   assume "finite B"
       
    81   thus "!!A. A \<subseteq> B ==> finite A"
       
    82   proof induct
       
    83     case empty
       
    84     thus ?case by simp
       
    85   next
       
    86     case (insert F x A)
       
    87     have A: "A \<subseteq> insert x F" and r: "A - {x} \<subseteq> F ==> finite (A - {x})" .
       
    88     show "finite A"
       
    89     proof cases
       
    90       assume x: "x \<in> A"
       
    91       with A have "A - {x} \<subseteq> F" by (simp add: subset_insert_iff)
       
    92       with r have "finite (A - {x})" .
       
    93       hence "finite (insert x (A - {x}))" ..
       
    94       also have "insert x (A - {x}) = A" by (rule insert_Diff)
       
    95       finally show ?thesis .
       
    96     next
       
    97       show "A \<subseteq> F ==> ?thesis" .
       
    98       assume "x \<notin> A"
       
    99       with A show "A \<subseteq> F" by (simp add: subset_insert_iff)
       
   100     qed
       
   101   qed
       
   102 qed
       
   103 
       
   104 lemma finite_Un [iff]: "finite (F Un G) = (finite F & finite G)"
       
   105   by (blast intro: finite_subset [of _ "X Un Y", standard] finite_UnI)
       
   106 
       
   107 lemma finite_Int [simp, intro]: "finite F | finite G ==> finite (F Int G)"
       
   108   -- {* The converse obviously fails. *}
       
   109   by (blast intro: finite_subset)
       
   110 
       
   111 lemma finite_insert [simp]: "finite (insert a A) = finite A"
       
   112   apply (subst insert_is_Un)
       
   113   apply (simp only: finite_Un)
       
   114   apply blast
       
   115   done
       
   116 
       
   117 lemma finite_imageI: "finite F ==> finite (h ` F)"
       
   118   -- {* The image of a finite set is finite. *}
       
   119   by (induct set: Finites) simp_all
       
   120 
       
   121 lemma finite_range_imageI:
       
   122     "finite (range g) ==> finite (range (%x. f (g x)))"
       
   123   apply (drule finite_imageI)
       
   124   apply simp
       
   125   done
       
   126 
       
   127 lemma finite_empty_induct:
       
   128   "finite A ==>
       
   129   P A ==> (!!a A. finite A ==> a:A ==> P A ==> P (A - {a})) ==> P {}"
       
   130 proof -
       
   131   assume "finite A"
       
   132     and "P A" and "!!a A. finite A ==> a:A ==> P A ==> P (A - {a})"
       
   133   have "P (A - A)"
       
   134   proof -
       
   135     fix c b :: "'a set"
       
   136     presume c: "finite c" and b: "finite b"
       
   137       and P1: "P b" and P2: "!!x y. finite y ==> x \<in> y ==> P y ==> P (y - {x})"
       
   138     from c show "c \<subseteq> b ==> P (b - c)"
       
   139     proof induct
       
   140       case empty
       
   141       from P1 show ?case by simp
       
   142     next
       
   143       case (insert F x)
       
   144       have "P (b - F - {x})"
       
   145       proof (rule P2)
       
   146         from _ b show "finite (b - F)" by (rule finite_subset) blast
       
   147         from insert show "x \<in> b - F" by simp
       
   148         from insert show "P (b - F)" by simp
       
   149       qed
       
   150       also have "b - F - {x} = b - insert x F" by (rule Diff_insert [symmetric])
       
   151       finally show ?case .
       
   152     qed
       
   153   next
       
   154     show "A \<subseteq> A" ..
       
   155   qed
       
   156   thus "P {}" by simp
       
   157 qed
       
   158 
       
   159 lemma finite_Diff [simp]: "finite B ==> finite (B - Ba)"
       
   160   by (rule Diff_subset [THEN finite_subset])
       
   161 
       
   162 lemma finite_Diff_insert [iff]: "finite (A - insert a B) = finite (A - B)"
       
   163   apply (subst Diff_insert)
       
   164   apply (case_tac "a : A - B")
       
   165    apply (rule finite_insert [symmetric, THEN trans])
       
   166    apply (subst insert_Diff)
       
   167     apply simp_all
       
   168   done
       
   169 
       
   170 
       
   171 lemma finite_imageD: "finite (f`A) ==> inj_on f A ==> finite A"
       
   172 proof -
       
   173   have aux: "!!A. finite (A - {}) = finite A" by simp
       
   174   fix B :: "'a set"
       
   175   assume "finite B"
       
   176   thus "!!A. f`A = B ==> inj_on f A ==> finite A"
       
   177     apply induct
       
   178      apply simp
       
   179     apply (subgoal_tac "EX y:A. f y = x & F = f ` (A - {y})")
       
   180      apply clarify
       
   181      apply (simp (no_asm_use) add: inj_on_def)
       
   182      apply (blast dest!: aux [THEN iffD1])
       
   183     apply atomize
       
   184     apply (erule_tac V = "ALL A. ?PP (A)" in thin_rl)
       
   185     apply (frule subsetD [OF equalityD2 insertI1])
       
   186     apply clarify
       
   187     apply (rule_tac x = xa in bexI)
       
   188      apply (simp_all add: inj_on_image_set_diff)
       
   189     done
       
   190 qed (rule refl)
       
   191 
       
   192 
       
   193 subsubsection {* The finite UNION of finite sets *}
       
   194 
       
   195 lemma finite_UN_I: "finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (UN a:A. B a)"
       
   196   by (induct set: Finites) simp_all
       
   197 
       
   198 text {*
       
   199   Strengthen RHS to
       
   200   @{prop "((ALL x:A. finite (B x)) & finite {x. x:A & B x ~= {}})"}?
       
   201 
       
   202   We'd need to prove
       
   203   @{prop "finite C ==> ALL A B. (UNION A B) <= C --> finite {x. x:A & B x ~= {}}"}
       
   204   by induction. *}
       
   205 
       
   206 lemma finite_UN [simp]: "finite A ==> finite (UNION A B) = (ALL x:A. finite (B x))"
       
   207   by (blast intro: finite_UN_I finite_subset)
       
   208 
       
   209 
       
   210 subsubsection {* Sigma of finite sets *}
       
   211 
       
   212 lemma finite_SigmaI [simp]:
       
   213     "finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (SIGMA a:A. B a)"
       
   214   by (unfold Sigma_def) (blast intro!: finite_UN_I)
       
   215 
       
   216 lemma finite_Prod_UNIV:
       
   217     "finite (UNIV::'a set) ==> finite (UNIV::'b set) ==> finite (UNIV::('a * 'b) set)"
       
   218   apply (subgoal_tac "(UNIV:: ('a * 'b) set) = Sigma UNIV (%x. UNIV)")
       
   219    apply (erule ssubst)
       
   220    apply (erule finite_SigmaI)
       
   221    apply auto
       
   222   done
       
   223 
       
   224 instance unit :: finite
       
   225 proof
       
   226   have "finite {()}" by simp
       
   227   also have "{()} = UNIV" by auto
       
   228   finally show "finite (UNIV :: unit set)" .
       
   229 qed
       
   230 
       
   231 instance * :: (finite, finite) finite
       
   232 proof
       
   233   show "finite (UNIV :: ('a \<times> 'b) set)"
       
   234   proof (rule finite_Prod_UNIV)
       
   235     show "finite (UNIV :: 'a set)" by (rule finite)
       
   236     show "finite (UNIV :: 'b set)" by (rule finite)
       
   237   qed
       
   238 qed
       
   239 
       
   240 
       
   241 subsubsection {* The powerset of a finite set *}
       
   242 
       
   243 lemma finite_Pow_iff [iff]: "finite (Pow A) = finite A"
       
   244 proof
       
   245   assume "finite (Pow A)"
       
   246   with _ have "finite ((%x. {x}) ` A)" by (rule finite_subset) blast
       
   247   thus "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp
       
   248 next
       
   249   assume "finite A"
       
   250   thus "finite (Pow A)"
       
   251     by induct (simp_all add: finite_UnI finite_imageI Pow_insert)
       
   252 qed
       
   253 
       
   254 lemma finite_converse [iff]: "finite (r^-1) = finite r"
       
   255   apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
       
   256    apply simp
       
   257    apply (rule iffI)
       
   258     apply (erule finite_imageD [unfolded inj_on_def])
       
   259     apply (simp split add: split_split)
       
   260    apply (erule finite_imageI)
       
   261   apply (simp add: converse_def image_def)
       
   262   apply auto
       
   263   apply (rule bexI)
       
   264    prefer 2 apply assumption
       
   265   apply simp
       
   266   done
       
   267 
       
   268 lemma finite_lessThan [iff]: (fixes k :: nat) "finite {..k(}"
       
   269   by (induct k) (simp_all add: lessThan_Suc)
       
   270 
       
   271 lemma finite_atMost [iff]: (fixes k :: nat) "finite {..k}"
       
   272   by (induct k) (simp_all add: atMost_Suc)
       
   273 
       
   274 lemma bounded_nat_set_is_finite:
       
   275     "(ALL i:N. i < (n::nat)) ==> finite N"
       
   276   -- {* A bounded set of natural numbers is finite. *}
       
   277   apply (rule finite_subset)
       
   278    apply (rule_tac [2] finite_lessThan)
       
   279   apply auto
       
   280   done
       
   281 
       
   282 
       
   283 subsubsection {* Finiteness of transitive closure *}
       
   284 
       
   285 text {* (Thanks to Sidi Ehmety) *}
       
   286 
       
   287 lemma finite_Field: "finite r ==> finite (Field r)"
       
   288   -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
       
   289   apply (induct set: Finites)
       
   290    apply (auto simp add: Field_def Domain_insert Range_insert)
       
   291   done
       
   292 
       
   293 lemma trancl_subset_Field2: "r^+ <= Field r \<times> Field r"
       
   294   apply clarify
       
   295   apply (erule trancl_induct)
       
   296    apply (auto simp add: Field_def)
       
   297   done
       
   298 
       
   299 lemma finite_trancl: "finite (r^+) = finite r"
       
   300   apply auto
       
   301    prefer 2
       
   302    apply (rule trancl_subset_Field2 [THEN finite_subset])
       
   303    apply (rule finite_SigmaI)
       
   304     prefer 3
       
   305     apply (blast intro: r_into_trancl finite_subset)
       
   306    apply (auto simp add: finite_Field)
       
   307   done
       
   308 
       
   309 
       
   310 subsection {* Finite cardinality *}
       
   311 
       
   312 text {*
       
   313   This definition, although traditional, is ugly to work with: @{text
       
   314   "card A == LEAST n. EX f. A = {f i | i. i < n}"}.  Therefore we have
       
   315   switched to an inductive one:
       
   316 *}
       
   317 
       
   318 consts cardR :: "('a set \<times> nat) set"
       
   319 
       
   320 inductive cardR
       
   321   intros
       
   322     EmptyI: "({}, 0) : cardR"
       
   323     InsertI: "(A, n) : cardR ==> a \<notin> A ==> (insert a A, Suc n) : cardR"
       
   324 
       
   325 constdefs
       
   326   card :: "'a set => nat"
       
   327   "card A == THE n. (A, n) : cardR"
       
   328 
       
   329 inductive_cases cardR_emptyE: "({}, n) : cardR"
       
   330 inductive_cases cardR_insertE: "(insert a A,n) : cardR"
       
   331 
       
   332 lemma cardR_SucD: "(A, n) : cardR ==> a : A --> (EX m. n = Suc m)"
       
   333   by (induct set: cardR) simp_all
       
   334 
       
   335 lemma cardR_determ_aux1:
       
   336     "(A, m): cardR ==> (!!n a. m = Suc n ==> a:A ==> (A - {a}, n) : cardR)"
       
   337   apply (induct set: cardR)
       
   338    apply auto
       
   339   apply (simp add: insert_Diff_if)
       
   340   apply auto
       
   341   apply (drule cardR_SucD)
       
   342   apply (blast intro!: cardR.intros)
       
   343   done
       
   344 
       
   345 lemma cardR_determ_aux2: "(insert a A, Suc m) : cardR ==> a \<notin> A ==> (A, m) : cardR"
       
   346   by (drule cardR_determ_aux1) auto
       
   347 
       
   348 lemma cardR_determ: "(A, m): cardR ==> (!!n. (A, n) : cardR ==> n = m)"
       
   349   apply (induct set: cardR)
       
   350    apply (safe elim!: cardR_emptyE cardR_insertE)
       
   351   apply (rename_tac B b m)
       
   352   apply (case_tac "a = b")
       
   353    apply (subgoal_tac "A = B")
       
   354     prefer 2 apply (blast elim: equalityE)
       
   355    apply blast
       
   356   apply (subgoal_tac "EX C. A = insert b C & B = insert a C")
       
   357    prefer 2
       
   358    apply (rule_tac x = "A Int B" in exI)
       
   359    apply (blast elim: equalityE)
       
   360   apply (frule_tac A = B in cardR_SucD)
       
   361   apply (blast intro!: cardR.intros dest!: cardR_determ_aux2)
       
   362   done
       
   363 
       
   364 lemma cardR_imp_finite: "(A, n) : cardR ==> finite A"
       
   365   by (induct set: cardR) simp_all
       
   366 
       
   367 lemma finite_imp_cardR: "finite A ==> EX n. (A, n) : cardR"
       
   368   by (induct set: Finites) (auto intro!: cardR.intros)
       
   369 
       
   370 lemma card_equality: "(A,n) : cardR ==> card A = n"
       
   371   by (unfold card_def) (blast intro: cardR_determ)
       
   372 
       
   373 lemma card_empty [simp]: "card {} = 0"
       
   374   by (unfold card_def) (blast intro!: cardR.intros elim!: cardR_emptyE)
       
   375 
       
   376 lemma card_insert_disjoint [simp]:
       
   377   "finite A ==> x \<notin> A ==> card (insert x A) = Suc(card A)"
       
   378 proof -
       
   379   assume x: "x \<notin> A"
       
   380   hence aux: "!!n. ((insert x A, n) : cardR) = (EX m. (A, m) : cardR & n = Suc m)"
       
   381     apply (auto intro!: cardR.intros)
       
   382     apply (rule_tac A1 = A in finite_imp_cardR [THEN exE])
       
   383      apply (force dest: cardR_imp_finite)
       
   384     apply (blast intro!: cardR.intros intro: cardR_determ)
       
   385     done
       
   386   assume "finite A"
       
   387   thus ?thesis
       
   388     apply (simp add: card_def aux)
       
   389     apply (rule the_equality)
       
   390      apply (auto intro: finite_imp_cardR
       
   391        cong: conj_cong simp: card_def [symmetric] card_equality)
       
   392     done
       
   393 qed
       
   394 
       
   395 lemma card_0_eq [simp]: "finite A ==> (card A = 0) = (A = {})"
       
   396   apply auto
       
   397   apply (drule_tac a = x in mk_disjoint_insert)
       
   398   apply clarify
       
   399   apply (rotate_tac -1)
       
   400   apply auto
       
   401   done
       
   402 
       
   403 lemma card_insert_if:
       
   404     "finite A ==> card (insert x A) = (if x:A then card A else Suc(card(A)))"
       
   405   by (simp add: insert_absorb)
       
   406 
       
   407 lemma card_Suc_Diff1: "finite A ==> x: A ==> Suc (card (A - {x})) = card A"
       
   408   apply (rule_tac t = A in insert_Diff [THEN subst])
       
   409    apply assumption
       
   410   apply simp
       
   411   done
       
   412 
       
   413 lemma card_Diff_singleton:
       
   414     "finite A ==> x: A ==> card (A - {x}) = card A - 1"
       
   415   by (simp add: card_Suc_Diff1 [symmetric])
       
   416 
       
   417 lemma card_Diff_singleton_if:
       
   418     "finite A ==> card (A-{x}) = (if x : A then card A - 1 else card A)"
       
   419   by (simp add: card_Diff_singleton)
       
   420 
       
   421 lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))"
       
   422   by (simp add: card_insert_if card_Suc_Diff1)
       
   423 
       
   424 lemma card_insert_le: "finite A ==> card A <= card (insert x A)"
       
   425   by (simp add: card_insert_if)
       
   426 
       
   427 lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)"
       
   428   apply (induct set: Finites)
       
   429    apply simp
       
   430   apply clarify
       
   431   apply (subgoal_tac "finite A & A - {x} <= F")
       
   432    prefer 2 apply (blast intro: finite_subset)
       
   433   apply atomize
       
   434   apply (drule_tac x = "A - {x}" in spec)
       
   435   apply (simp add: card_Diff_singleton_if split add: split_if_asm)
       
   436   apply (case_tac "card A")
       
   437    apply auto
       
   438   done
       
   439 
       
   440 lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B"
       
   441   apply (simp add: psubset_def linorder_not_le [symmetric])
       
   442   apply (blast dest: card_seteq)
       
   443   done
       
   444 
       
   445 lemma card_mono: "finite B ==> A <= B ==> card A <= card B"
       
   446   apply (case_tac "A = B")
       
   447    apply simp
       
   448   apply (simp add: linorder_not_less [symmetric])
       
   449   apply (blast dest: card_seteq intro: order_less_imp_le)
       
   450   done
       
   451 
       
   452 lemma card_Un_Int: "finite A ==> finite B
       
   453     ==> card A + card B = card (A Un B) + card (A Int B)"
       
   454   apply (induct set: Finites)
       
   455    apply simp
       
   456   apply (simp add: insert_absorb Int_insert_left)
       
   457   done
       
   458 
       
   459 lemma card_Un_disjoint: "finite A ==> finite B
       
   460     ==> A Int B = {} ==> card (A Un B) = card A + card B"
       
   461   by (simp add: card_Un_Int)
       
   462 
       
   463 lemma card_Diff_subset:
       
   464     "finite A ==> B <= A ==> card A - card B = card (A - B)"
       
   465   apply (subgoal_tac "(A - B) Un B = A")
       
   466    prefer 2 apply blast
       
   467   apply (rule add_right_cancel [THEN iffD1])
       
   468   apply (rule card_Un_disjoint [THEN subst])
       
   469      apply (erule_tac [4] ssubst)
       
   470      prefer 3 apply blast
       
   471     apply (simp_all add: add_commute not_less_iff_le
       
   472       add_diff_inverse card_mono finite_subset)
       
   473   done
       
   474 
       
   475 lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A"
       
   476   apply (rule Suc_less_SucD)
       
   477   apply (simp add: card_Suc_Diff1)
       
   478   done
       
   479 
       
   480 lemma card_Diff2_less:
       
   481     "finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A"
       
   482   apply (case_tac "x = y")
       
   483    apply (simp add: card_Diff1_less)
       
   484   apply (rule less_trans)
       
   485    prefer 2 apply (auto intro!: card_Diff1_less)
       
   486   done
       
   487 
       
   488 lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A"
       
   489   apply (case_tac "x : A")
       
   490    apply (simp_all add: card_Diff1_less less_imp_le)
       
   491   done
       
   492 
       
   493 lemma card_psubset: "finite B ==> A \<subseteq> B ==> card A < card B ==> A < B"
       
   494   apply (erule psubsetI)
       
   495   apply blast
       
   496   done
       
   497 
       
   498 
       
   499 subsubsection {* Cardinality of image *}
       
   500 
       
   501 lemma card_image_le: "finite A ==> card (f ` A) <= card A"
       
   502   apply (induct set: Finites)
       
   503    apply simp
       
   504   apply (simp add: le_SucI finite_imageI card_insert_if)
       
   505   done
       
   506 
       
   507 lemma card_image: "finite A ==> inj_on f A ==> card (f ` A) = card A"
       
   508   apply (induct set: Finites)
       
   509    apply simp_all
       
   510   apply atomize
       
   511   apply safe
       
   512    apply (unfold inj_on_def)
       
   513    apply blast
       
   514   apply (subst card_insert_disjoint)
       
   515     apply (erule finite_imageI)
       
   516    apply blast
       
   517   apply blast
       
   518   done
       
   519 
       
   520 lemma endo_inj_surj: "finite A ==> f ` A \<subseteq> A ==> inj_on f A ==> f ` A = A"
       
   521   by (simp add: card_seteq card_image)
       
   522 
       
   523 
       
   524 subsubsection {* Cardinality of the Powerset *}
       
   525 
       
   526 lemma card_Pow: "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A"  (* FIXME numeral 2 (!?) *)
       
   527   apply (induct set: Finites)
       
   528    apply (simp_all add: Pow_insert)
       
   529   apply (subst card_Un_disjoint)
       
   530      apply blast
       
   531     apply (blast intro: finite_imageI)
       
   532    apply blast
       
   533   apply (subgoal_tac "inj_on (insert x) (Pow F)")
       
   534    apply (simp add: card_image Pow_insert)
       
   535   apply (unfold inj_on_def)
       
   536   apply (blast elim!: equalityE)
       
   537   done
       
   538 
       
   539 text {*
       
   540   \medskip Relates to equivalence classes.  Based on a theorem of
       
   541   F. Kammüller's.  The @{prop "finite C"} premise is redundant.
       
   542 *}
       
   543 
       
   544 lemma dvd_partition:
       
   545   "finite C ==> finite (Union C) ==>
       
   546     ALL c : C. k dvd card c ==>
       
   547     (ALL c1: C. ALL c2: C. c1 ~= c2 --> c1 Int c2 = {}) ==>
       
   548   k dvd card (Union C)"
       
   549   apply (induct set: Finites)
       
   550    apply simp_all
       
   551   apply clarify
       
   552   apply (subst card_Un_disjoint)
       
   553   apply (auto simp add: dvd_add disjoint_eq_subset_Compl)
       
   554   done
       
   555 
       
   556 
       
   557 subsection {* A fold functional for finite sets *}
       
   558 
       
   559 text {*
       
   560   For @{text n} non-negative we have @{text "fold f e {x1, ..., xn} =
       
   561   f x1 (... (f xn e))"} where @{text f} is at least left-commutative.
       
   562 *}
       
   563 
       
   564 consts
       
   565   foldSet :: "('b => 'a => 'a) => 'a => ('b set \<times> 'a) set"
       
   566 
       
   567 inductive "foldSet f e"
       
   568   intros
       
   569     emptyI [intro]: "({}, e) : foldSet f e"
       
   570     insertI [intro]: "x \<notin> A ==> (A, y) : foldSet f e ==> (insert x A, f x y) : foldSet f e"
       
   571 
       
   572 inductive_cases empty_foldSetE [elim!]: "({}, x) : foldSet f e"
       
   573 
       
   574 constdefs
       
   575   fold :: "('b => 'a => 'a) => 'a => 'b set => 'a"
       
   576   "fold f e A == THE x. (A, x) : foldSet f e"
       
   577 
       
   578 lemma Diff1_foldSet: "(A - {x}, y) : foldSet f e ==> x: A ==> (A, f x y) : foldSet f e"
       
   579   apply (erule insert_Diff [THEN subst], rule foldSet.intros)
       
   580    apply auto
       
   581   done
       
   582 
       
   583 lemma foldSet_imp_finite [simp]: "(A, x) : foldSet f e ==> finite A"
       
   584   by (induct set: foldSet) auto
       
   585 
       
   586 lemma finite_imp_foldSet: "finite A ==> EX x. (A, x) : foldSet f e"
       
   587   by (induct set: Finites) auto
       
   588 
       
   589 
       
   590 subsubsection {* Left-commutative operations *}
       
   591 
       
   592 locale LC =
       
   593   fixes f :: "'b => 'a => 'a"    (infixl "\<cdot>" 70)
       
   594   assumes left_commute: "x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)"
       
   595 
       
   596 lemma (in LC) foldSet_determ_aux:
       
   597   "ALL A x. card A < n --> (A, x) : foldSet f e -->
       
   598     (ALL y. (A, y) : foldSet f e --> y = x)"
       
   599   apply (induct n)
       
   600    apply (auto simp add: less_Suc_eq)
       
   601   apply (erule foldSet.cases)
       
   602    apply blast
       
   603   apply (erule foldSet.cases)
       
   604    apply blast
       
   605   apply clarify
       
   606   txt {* force simplification of @{text "card A < card (insert ...)"}. *}
       
   607   apply (erule rev_mp)
       
   608   apply (simp add: less_Suc_eq_le)
       
   609   apply (rule impI)
       
   610   apply (rename_tac Aa xa ya Ab xb yb, case_tac "xa = xb")
       
   611    apply (subgoal_tac "Aa = Ab")
       
   612     prefer 2 apply (blast elim!: equalityE)
       
   613    apply blast
       
   614   txt {* case @{prop "xa \<notin> xb"}. *}
       
   615   apply (subgoal_tac "Aa - {xb} = Ab - {xa} & xb : Aa & xa : Ab")
       
   616    prefer 2 apply (blast elim!: equalityE)
       
   617   apply clarify
       
   618   apply (subgoal_tac "Aa = insert xb Ab - {xa}")
       
   619    prefer 2 apply blast
       
   620   apply (subgoal_tac "card Aa <= card Ab")
       
   621    prefer 2
       
   622    apply (rule Suc_le_mono [THEN subst])
       
   623    apply (simp add: card_Suc_Diff1)
       
   624   apply (rule_tac A1 = "Aa - {xb}" and f1 = f in finite_imp_foldSet [THEN exE])
       
   625   apply (blast intro: foldSet_imp_finite finite_Diff)
       
   626   apply (frule (1) Diff1_foldSet)
       
   627   apply (subgoal_tac "ya = f xb x")
       
   628    prefer 2 apply (blast del: equalityCE)
       
   629   apply (subgoal_tac "(Ab - {xa}, x) : foldSet f e")
       
   630    prefer 2 apply simp
       
   631   apply (subgoal_tac "yb = f xa x")
       
   632    prefer 2 apply (blast del: equalityCE dest: Diff1_foldSet)
       
   633   apply (simp (no_asm_simp) add: left_commute)
       
   634   done
       
   635 
       
   636 lemma (in LC) foldSet_determ: "(A, x) : foldSet f e ==> (A, y) : foldSet f e ==> y = x"
       
   637   by (blast intro: foldSet_determ_aux [rule_format])
       
   638 
       
   639 lemma (in LC) fold_equality: "(A, y) : foldSet f e ==> fold f e A = y"
       
   640   by (unfold fold_def) (blast intro: foldSet_determ)
       
   641 
       
   642 lemma fold_empty [simp]: "fold f e {} = e"
       
   643   by (unfold fold_def) blast
       
   644 
       
   645 lemma (in LC) fold_insert_aux: "x \<notin> A ==>
       
   646     ((insert x A, v) : foldSet f e) =
       
   647     (EX y. (A, y) : foldSet f e & v = f x y)"
       
   648   apply auto
       
   649   apply (rule_tac A1 = A and f1 = f in finite_imp_foldSet [THEN exE])
       
   650    apply (fastsimp dest: foldSet_imp_finite)
       
   651   apply (blast intro: foldSet_determ)
       
   652   done
       
   653 
       
   654 lemma (in LC) fold_insert:
       
   655     "finite A ==> x \<notin> A ==> fold f e (insert x A) = f x (fold f e A)"
       
   656   apply (unfold fold_def)
       
   657   apply (simp add: fold_insert_aux)
       
   658   apply (rule the_equality)
       
   659   apply (auto intro: finite_imp_foldSet
       
   660     cong add: conj_cong simp add: fold_def [symmetric] fold_equality)
       
   661   done
       
   662 
       
   663 lemma (in LC) fold_commute: "finite A ==> (!!e. f x (fold f e A) = fold f (f x e) A)"
       
   664   apply (induct set: Finites)
       
   665    apply simp
       
   666   apply (simp add: left_commute fold_insert)
       
   667   done
       
   668 
       
   669 lemma (in LC) fold_nest_Un_Int:
       
   670   "finite A ==> finite B
       
   671     ==> fold f (fold f e B) A = fold f (fold f e (A Int B)) (A Un B)"
       
   672   apply (induct set: Finites)
       
   673    apply simp
       
   674   apply (simp add: fold_insert fold_commute Int_insert_left insert_absorb)
       
   675   done
       
   676 
       
   677 lemma (in LC) fold_nest_Un_disjoint:
       
   678   "finite A ==> finite B ==> A Int B = {}
       
   679     ==> fold f e (A Un B) = fold f (fold f e B) A"
       
   680   by (simp add: fold_nest_Un_Int)
       
   681 
       
   682 declare foldSet_imp_finite [simp del]
       
   683     empty_foldSetE [rule del]  foldSet.intros [rule del]
       
   684   -- {* Delete rules to do with @{text foldSet} relation. *}
       
   685 
       
   686 
       
   687 
       
   688 subsubsection {* Commutative monoids *}
       
   689 
       
   690 text {*
       
   691   We enter a more restrictive context, with @{text "f :: 'a => 'a => 'a"}
       
   692   instead of @{text "'b => 'a => 'a"}.
       
   693 *}
       
   694 
       
   695 locale ACe =
       
   696   fixes f :: "'a => 'a => 'a"    (infixl "\<cdot>" 70)
       
   697     and e :: 'a
       
   698   assumes ident [simp]: "x \<cdot> e = x"
       
   699     and commute: "x \<cdot> y = y \<cdot> x"
       
   700     and assoc: "(x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)"
       
   701 
       
   702 lemma (in ACe) left_commute: "x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)"
       
   703 proof -
       
   704   have "x \<cdot> (y \<cdot> z) = (y \<cdot> z) \<cdot> x" by (simp only: commute)
       
   705   also have "... = y \<cdot> (z \<cdot> x)" by (simp only: assoc)
       
   706   also have "z \<cdot> x = x \<cdot> z" by (simp only: commute)
       
   707   finally show ?thesis .
       
   708 qed
       
   709 
       
   710 lemma (in ACe)
       
   711     AC: "(x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)"  "x \<cdot> y = y \<cdot> x"  "x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)"
       
   712   by (rule assoc, rule commute, rule left_commute)  (* FIXME localize "lemmas" (!??) *)
       
   713 
       
   714 lemma (in ACe [simp]) left_ident: "e \<cdot> x = x"
       
   715 proof -
       
   716   have "x \<cdot> e = x" by (rule ident)
       
   717   thus ?thesis by (subst commute)
       
   718 qed
       
   719 
       
   720 lemma (in ACe) fold_Un_Int:
       
   721   "finite A ==> finite B ==>
       
   722     fold f e A \<cdot> fold f e B = fold f e (A Un B) \<cdot> fold f e (A Int B)"
       
   723   apply (induct set: Finites)
       
   724    apply simp
       
   725   apply (simp add: AC fold_insert insert_absorb Int_insert_left)
       
   726   done
       
   727 
       
   728 lemma (in ACe) fold_Un_disjoint:
       
   729   "finite A ==> finite B ==> A Int B = {} ==>
       
   730     fold f e (A Un B) = fold f e A \<cdot> fold f e B"
       
   731   by (simp add: fold_Un_Int)
       
   732 
       
   733 lemma (in ACe) fold_Un_disjoint2:
       
   734   "finite A ==> finite B ==> A Int B = {} ==>
       
   735     fold (f o g) e (A Un B) = fold (f o g) e A \<cdot> fold (f o g) e B"
       
   736 proof -
       
   737   assume b: "finite B"
       
   738   assume "finite A"
       
   739   thus "A Int B = {} ==>
       
   740     fold (f o g) e (A Un B) = fold (f o g) e A \<cdot> fold (f o g) e B"
       
   741   proof induct
       
   742     case empty
       
   743     thus ?case by simp
       
   744   next
       
   745     case (insert F x)
       
   746     have "fold (f \<circ> g) e (insert x F \<union> B) = fold (f \<circ> g) e (insert x (F \<union> B))"
       
   747       by simp
       
   748     also have "... = (f \<circ> g) x (fold (f \<circ> g) e (F \<union> B))"
       
   749       by (rule fold_insert) (insert b insert, auto simp add: left_commute)  (* FIXME import of fold_insert (!?) *)
       
   750     also from insert have "fold (f \<circ> g) e (F \<union> B) =
       
   751       fold (f \<circ> g) e F \<cdot> fold (f \<circ> g) e B" by blast
       
   752     also have "(f \<circ> g) x ... = (f \<circ> g) x (fold (f \<circ> g) e F) \<cdot> fold (f \<circ> g) e B"
       
   753       by (simp add: AC)
       
   754     also have "(f \<circ> g) x (fold (f \<circ> g) e F) = fold (f \<circ> g) e (insert x F)"
       
   755       by (rule fold_insert [symmetric]) (insert b insert, auto simp add: left_commute)
       
   756     finally show ?case .
       
   757   qed
       
   758 qed
       
   759 
       
   760 
       
   761 subsection {* Generalized summation over a set *}
       
   762 
       
   763 constdefs
       
   764   setsum :: "('a => 'b) => 'a set => 'b::plus_ac0"
       
   765   "setsum f A == if finite A then fold (op + o f) 0 A else 0"
       
   766 
       
   767 syntax
       
   768   "_setsum" :: "idt => 'a set => 'b => 'b::plus_ac0"    ("\<Sum>_:_. _" [0, 51, 10] 10)
       
   769 syntax (xsymbols)
       
   770   "_setsum" :: "idt => 'a set => 'b => 'b::plus_ac0"    ("\<Sum>_\<in>_. _" [0, 51, 10] 10)
       
   771 translations
       
   772   "\<Sum>i:A. b" == "setsum (%i. b) A"  -- {* Beware of argument permutation! *}
       
   773 
       
   774 
       
   775 lemma setsum_empty [simp]: "setsum f {} = 0"
       
   776   by (simp add: setsum_def)
       
   777 
       
   778 lemma setsum_insert [simp]:
       
   779     "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"
       
   780   by (simp add: setsum_def fold_insert plus_ac0_left_commute)
       
   781 
       
   782 lemma setsum_0: "setsum (\<lambda>i. 0) A = 0"
       
   783   apply (case_tac "finite A")
       
   784    prefer 2 apply (simp add: setsum_def)
       
   785   apply (erule finite_induct)
       
   786    apply auto
       
   787   done
       
   788 
       
   789 lemma setsum_eq_0_iff [simp]:
       
   790     "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))"
       
   791   by (induct set: Finites) auto
       
   792 
       
   793 lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a"
       
   794   apply (case_tac "finite A")
       
   795    prefer 2 apply (simp add: setsum_def)
       
   796   apply (erule rev_mp)
       
   797   apply (erule finite_induct)
       
   798    apply auto
       
   799   done
       
   800 
       
   801 lemma card_eq_setsum: "finite A ==> card A = setsum (\<lambda>x. 1) A"
       
   802   -- {* Could allow many @{text "card"} proofs to be simplified. *}
       
   803   by (induct set: Finites) auto
       
   804 
       
   805 lemma setsum_Un_Int: "finite A ==> finite B
       
   806     ==> setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"
       
   807   -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
       
   808   apply (induct set: Finites)
       
   809    apply simp
       
   810   apply (simp add: plus_ac0 Int_insert_left insert_absorb)
       
   811   done
       
   812 
       
   813 lemma setsum_Un_disjoint: "finite A ==> finite B
       
   814   ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
       
   815   apply (subst setsum_Un_Int [symmetric])
       
   816     apply auto
       
   817   done
       
   818 
       
   819 lemma setsum_UN_disjoint: (fixes f :: "'a => 'b::plus_ac0")
       
   820   "finite I ==> (ALL i:I. finite (A i)) ==>
       
   821       (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
       
   822     setsum f (UNION I A) = setsum (\<lambda>i. setsum f (A i)) I"
       
   823   apply (induct set: Finites)
       
   824    apply simp
       
   825   apply atomize
       
   826   apply (subgoal_tac "ALL i:F. x \<noteq> i")
       
   827    prefer 2 apply blast
       
   828   apply (subgoal_tac "A x Int UNION F A = {}")
       
   829    prefer 2 apply blast
       
   830   apply (simp add: setsum_Un_disjoint)
       
   831   done
       
   832 
       
   833 lemma setsum_addf: "setsum (\<lambda>x. f x + g x) A = (setsum f A + setsum g A)"
       
   834   apply (case_tac "finite A")
       
   835    prefer 2 apply (simp add: setsum_def)
       
   836   apply (erule finite_induct)
       
   837    apply auto
       
   838   apply (simp add: plus_ac0)
       
   839   done
       
   840 
       
   841 lemma setsum_Un: "finite A ==> finite B ==>
       
   842     (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
       
   843   -- {* For the natural numbers, we have subtraction. *}
       
   844   apply (subst setsum_Un_Int [symmetric])
       
   845     apply auto
       
   846   done
       
   847 
       
   848 lemma setsum_diff1: "(setsum f (A - {a}) :: nat) =
       
   849     (if a:A then setsum f A - f a else setsum f A)"
       
   850   apply (case_tac "finite A")
       
   851    prefer 2 apply (simp add: setsum_def)
       
   852   apply (erule finite_induct)
       
   853    apply (auto simp add: insert_Diff_if)
       
   854   apply (drule_tac a = a in mk_disjoint_insert)
       
   855   apply auto
       
   856   done
       
   857 
       
   858 lemma setsum_cong:
       
   859   "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B"
       
   860   apply (case_tac "finite B")
       
   861    prefer 2 apply (simp add: setsum_def)
       
   862   apply simp
       
   863   apply (subgoal_tac "ALL C. C <= B --> (ALL x:C. f x = g x) --> setsum f C = setsum g C")
       
   864    apply simp
       
   865   apply (erule finite_induct)
       
   866   apply simp
       
   867   apply (simp add: subset_insert_iff)
       
   868   apply clarify
       
   869   apply (subgoal_tac "finite C")
       
   870    prefer 2 apply (blast dest: finite_subset [COMP swap_prems_rl])
       
   871   apply (subgoal_tac "C = insert x (C - {x})")
       
   872    prefer 2 apply blast
       
   873   apply (erule ssubst)
       
   874   apply (drule spec)
       
   875   apply (erule (1) notE impE)
       
   876   apply (simp add: Ball_def)
       
   877   done
       
   878 
       
   879 
       
   880 text {*
       
   881   \medskip Basic theorem about @{text "choose"}.  By Florian
       
   882   Kammüller, tidied by LCP.
       
   883 *}
       
   884 
       
   885 lemma card_s_0_eq_empty:
       
   886     "finite A ==> card {B. B \<subseteq> A & card B = 0} = 1"
       
   887   apply (simp cong add: conj_cong add: finite_subset [THEN card_0_eq])
       
   888   apply (simp cong add: rev_conj_cong)
       
   889   done
       
   890 
       
   891 lemma choose_deconstruct: "finite M ==> x \<notin> M
       
   892   ==> {s. s <= insert x M & card(s) = Suc k}
       
   893        = {s. s <= M & card(s) = Suc k} Un
       
   894          {s. EX t. t <= M & card(t) = k & s = insert x t}"
       
   895   apply safe
       
   896    apply (auto intro: finite_subset [THEN card_insert_disjoint])
       
   897   apply (drule_tac x = "xa - {x}" in spec)
       
   898   apply (subgoal_tac "x ~: xa")
       
   899    apply auto
       
   900   apply (erule rev_mp, subst card_Diff_singleton)
       
   901   apply (auto intro: finite_subset)
       
   902   done
       
   903 
       
   904 lemma card_inj_on_le:
       
   905     "finite A ==> finite B ==> f ` A \<subseteq> B ==> inj_on f A ==> card A <= card B"
       
   906   by (auto intro: card_mono simp add: card_image [symmetric])
       
   907 
       
   908 lemma card_bij_eq: "finite A ==> finite B ==>
       
   909   f ` A \<subseteq> B ==> inj_on f A ==> g ` B \<subseteq> A ==> inj_on g B ==> card A = card B"
       
   910   by (auto intro: le_anti_sym card_inj_on_le)
       
   911 
       
   912 lemma constr_bij: "finite A ==> x \<notin> A ==>
       
   913   card {B. EX C. C <= A & card(C) = k & B = insert x C} =
       
   914     card {B. B <= A & card(B) = k}"
       
   915   apply (rule_tac f = "%s. s - {x}" and g = "insert x" in card_bij_eq)
       
   916        apply (rule_tac B = "Pow (insert x A) " in finite_subset)
       
   917         apply (rule_tac [3] B = "Pow (A) " in finite_subset)
       
   918          apply fast+
       
   919      txt {* arity *}
       
   920      apply (auto elim!: equalityE simp add: inj_on_def)
       
   921   apply (subst Diff_insert0)
       
   922   apply auto
       
   923   done
       
   924 
       
   925 text {*
       
   926   Main theorem: combinatorial statement about number of subsets of a set.
       
   927 *}
       
   928 
       
   929 lemma n_sub_lemma:
       
   930   "!!A. finite A ==> card {B. B <= A & card B = k} = (card A choose k)"
       
   931   apply (induct k)
       
   932    apply (simp add: card_s_0_eq_empty)
       
   933   apply atomize
       
   934   apply (rotate_tac -1, erule finite_induct)
       
   935    apply (simp_all (no_asm_simp) cong add: conj_cong add: card_s_0_eq_empty choose_deconstruct)
       
   936   apply (subst card_Un_disjoint)
       
   937      prefer 4 apply (force simp add: constr_bij)
       
   938     prefer 3 apply force
       
   939    prefer 2 apply (blast intro: finite_Pow_iff [THEN iffD2]
       
   940      finite_subset [of _ "Pow (insert x F)", standard])
       
   941   apply (blast intro: finite_Pow_iff [THEN iffD2, THEN [2] finite_subset])
       
   942   done
       
   943 
       
   944 theorem n_subsets: "finite A ==> card {B. B <= A & card(B) = k} = (card A choose k)"
       
   945   by (simp add: n_sub_lemma)
       
   946 
       
   947 end