src/HOL/Finite_Set.thy
changeset 63404 a95e7432d86c
parent 63365 5340fb6633d0
child 63561 fba08009ff3e
equal deleted inserted replaced
63403:a962f349c8c9 63404:a95e7432d86c
    14 context
    14 context
    15   notes [[inductive_internals]]
    15   notes [[inductive_internals]]
    16 begin
    16 begin
    17 
    17 
    18 inductive finite :: "'a set \<Rightarrow> bool"
    18 inductive finite :: "'a set \<Rightarrow> bool"
    19   where
    19 where
    20     emptyI [simp, intro!]: "finite {}"
    20   emptyI [simp, intro!]: "finite {}"
    21   | insertI [simp, intro!]: "finite A \<Longrightarrow> finite (insert a A)"
    21 | insertI [simp, intro!]: "finite A \<Longrightarrow> finite (insert a A)"
    22 
    22 
    23 end
    23 end
    24 
    24 
    25 simproc_setup finite_Collect ("finite (Collect P)") = \<open>K Set_Comprehension_Pointfree.simproc\<close>
    25 simproc_setup finite_Collect ("finite (Collect P)") = \<open>K Set_Comprehension_Pointfree.simproc\<close>
    26 
    26 
    30   \<comment> \<open>Discharging \<open>x \<notin> F\<close> entails extra work.\<close>
    30   \<comment> \<open>Discharging \<open>x \<notin> F\<close> entails extra work.\<close>
    31   assumes "finite F"
    31   assumes "finite F"
    32   assumes "P {}"
    32   assumes "P {}"
    33     and insert: "\<And>x F. finite F \<Longrightarrow> x \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert x F)"
    33     and insert: "\<And>x F. finite F \<Longrightarrow> x \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert x F)"
    34   shows "P F"
    34   shows "P F"
    35 using \<open>finite F\<close>
    35   using \<open>finite F\<close>
    36 proof induct
    36 proof induct
    37   show "P {}" by fact
    37   show "P {}" by fact
    38   fix x F assume F: "finite F" and P: "P F"
    38 next
       
    39   fix x F
       
    40   assume F: "finite F" and P: "P F"
    39   show "P (insert x F)"
    41   show "P (insert x F)"
    40   proof cases
    42   proof cases
    41     assume "x \<in> F"
    43     assume "x \<in> F"
    42     hence "insert x F = F" by (rule insert_absorb)
    44     then have "insert x F = F" by (rule insert_absorb)
    43     with P show ?thesis by (simp only:)
    45     with P show ?thesis by (simp only:)
    44   next
    46   next
    45     assume "x \<notin> F"
    47     assume "x \<notin> F"
    46     from F this P show ?thesis by (rule insert)
    48     from F this P show ?thesis by (rule insert)
    47   qed
    49   qed
    48 qed
    50 qed
    49 
    51 
    50 lemma infinite_finite_induct [case_names infinite empty insert]:
    52 lemma infinite_finite_induct [case_names infinite empty insert]:
    51   assumes infinite: "\<And>A. \<not> finite A \<Longrightarrow> P A"
    53   assumes infinite: "\<And>A. \<not> finite A \<Longrightarrow> P A"
    52   assumes empty: "P {}"
    54     and empty: "P {}"
    53   assumes insert: "\<And>x F. finite F \<Longrightarrow> x \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert x F)"
    55     and insert: "\<And>x F. finite F \<Longrightarrow> x \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert x F)"
    54   shows "P A"
    56   shows "P A"
    55 proof (cases "finite A")
    57 proof (cases "finite A")
    56   case False with infinite show ?thesis .
    58   case False
    57 next
    59   with infinite show ?thesis .
    58   case True then show ?thesis by (induct A) (fact empty insert)+
    60 next
       
    61   case True
       
    62   then show ?thesis by (induct A) (fact empty insert)+
    59 qed
    63 qed
    60 
    64 
    61 
    65 
    62 subsubsection \<open>Choice principles\<close>
    66 subsubsection \<open>Choice principles\<close>
    63 
    67 
    69   then show ?thesis by blast
    73   then show ?thesis by blast
    70 qed
    74 qed
    71 
    75 
    72 text \<open>A finite choice principle. Does not need the SOME choice operator.\<close>
    76 text \<open>A finite choice principle. Does not need the SOME choice operator.\<close>
    73 
    77 
    74 lemma finite_set_choice:
    78 lemma finite_set_choice: "finite A \<Longrightarrow> \<forall>x\<in>A. \<exists>y. P x y \<Longrightarrow> \<exists>f. \<forall>x\<in>A. P x (f x)"
    75   "finite A \<Longrightarrow> \<forall>x\<in>A. \<exists>y. P x y \<Longrightarrow> \<exists>f. \<forall>x\<in>A. P x (f x)"
       
    76 proof (induct rule: finite_induct)
    79 proof (induct rule: finite_induct)
    77   case empty then show ?case by simp
    80   case empty
       
    81   then show ?case by simp
    78 next
    82 next
    79   case (insert a A)
    83   case (insert a A)
    80   then obtain f b where f: "ALL x:A. P x (f x)" and ab: "P a b" by auto
    84   then obtain f b where f: "\<forall>x\<in>A. P x (f x)" and ab: "P a b"
    81   show ?case (is "EX f. ?P f")
    85     by auto
       
    86   show ?case (is "\<exists>f. ?P f")
    82   proof
    87   proof
    83     show "?P(%x. if x = a then b else f x)" using f ab by auto
    88     show "?P (\<lambda>x. if x = a then b else f x)"
       
    89       using f ab by auto
    84   qed
    90   qed
    85 qed
    91 qed
    86 
    92 
    87 
    93 
    88 subsubsection \<open>Finite sets are the images of initial segments of natural numbers\<close>
    94 subsubsection \<open>Finite sets are the images of initial segments of natural numbers\<close>
    89 
    95 
    90 lemma finite_imp_nat_seg_image_inj_on:
    96 lemma finite_imp_nat_seg_image_inj_on:
    91   assumes "finite A" 
    97   assumes "finite A"
    92   shows "\<exists>(n::nat) f. A = f ` {i. i < n} \<and> inj_on f {i. i < n}"
    98   shows "\<exists>(n::nat) f. A = f ` {i. i < n} \<and> inj_on f {i. i < n}"
    93 using assms
    99   using assms
    94 proof induct
   100 proof induct
    95   case empty
   101   case empty
    96   show ?case
   102   show ?case
    97   proof
   103   proof
    98     show "\<exists>f. {} = f ` {i::nat. i < 0} \<and> inj_on f {i. i < 0}" by simp 
   104     show "\<exists>f. {} = f ` {i::nat. i < 0} \<and> inj_on f {i. i < 0}"
       
   105       by simp
    99   qed
   106   qed
   100 next
   107 next
   101   case (insert a A)
   108   case (insert a A)
   102   have notinA: "a \<notin> A" by fact
   109   have notinA: "a \<notin> A" by fact
   103   from insert.hyps obtain n f
   110   from insert.hyps obtain n f where "A = f ` {i::nat. i < n}" "inj_on f {i. i < n}"
   104     where "A = f ` {i::nat. i < n}" "inj_on f {i. i < n}" by blast
   111     by blast
   105   hence "insert a A = f(n:=a) ` {i. i < Suc n}"
   112   then have "insert a A = f(n:=a) ` {i. i < Suc n}" and "inj_on (f(n:=a)) {i. i < Suc n}"
   106         "inj_on (f(n:=a)) {i. i < Suc n}" using notinA
   113     using notinA by (auto simp add: image_def Ball_def inj_on_def less_Suc_eq)
   107     by (auto simp add: image_def Ball_def inj_on_def less_Suc_eq)
   114   then show ?case by blast
   108   thus ?case by blast
   115 qed
   109 qed
   116 
   110 
   117 lemma nat_seg_image_imp_finite: "A = f ` {i::nat. i < n} \<Longrightarrow> finite A"
   111 lemma nat_seg_image_imp_finite:
       
   112   "A = f ` {i::nat. i < n} \<Longrightarrow> finite A"
       
   113 proof (induct n arbitrary: A)
   118 proof (induct n arbitrary: A)
   114   case 0 thus ?case by simp
   119   case 0
       
   120   then show ?case by simp
   115 next
   121 next
   116   case (Suc n)
   122   case (Suc n)
   117   let ?B = "f ` {i. i < n}"
   123   let ?B = "f ` {i. i < n}"
   118   have finB: "finite ?B" by(rule Suc.hyps[OF refl])
   124   have finB: "finite ?B" by (rule Suc.hyps[OF refl])
   119   show ?case
   125   show ?case
   120   proof cases
   126   proof (cases "\<exists>k<n. f n = f k")
   121     assume "\<exists>k<n. f n = f k"
   127     case True
   122     hence "A = ?B" using Suc.prems by(auto simp:less_Suc_eq)
   128     then have "A = ?B"
   123     thus ?thesis using finB by simp
   129       using Suc.prems by (auto simp:less_Suc_eq)
       
   130     then show ?thesis
       
   131       using finB by simp
   124   next
   132   next
   125     assume "\<not>(\<exists> k<n. f n = f k)"
   133     case False
   126     hence "A = insert (f n) ?B" using Suc.prems by(auto simp:less_Suc_eq)
   134     then have "A = insert (f n) ?B"
   127     thus ?thesis using finB by simp
   135       using Suc.prems by (auto simp:less_Suc_eq)
   128   qed
   136     then show ?thesis using finB by simp
   129 qed
   137   qed
   130 
   138 qed
   131 lemma finite_conv_nat_seg_image:
   139 
   132   "finite A \<longleftrightarrow> (\<exists>(n::nat) f. A = f ` {i::nat. i < n})"
   140 lemma finite_conv_nat_seg_image: "finite A \<longleftrightarrow> (\<exists>(n::nat) f. A = f ` {i::nat. i < n})"
   133   by (blast intro: nat_seg_image_imp_finite dest: finite_imp_nat_seg_image_inj_on)
   141   by (blast intro: nat_seg_image_imp_finite dest: finite_imp_nat_seg_image_inj_on)
   134 
   142 
   135 lemma finite_imp_inj_to_nat_seg:
   143 lemma finite_imp_inj_to_nat_seg:
   136   assumes "finite A"
   144   assumes "finite A"
   137   shows "\<exists>f n::nat. f ` A = {i. i < n} \<and> inj_on f A"
   145   shows "\<exists>f n::nat. f ` A = {i. i < n} \<and> inj_on f A"
   138 proof -
   146 proof -
   139   from finite_imp_nat_seg_image_inj_on[OF \<open>finite A\<close>]
   147   from finite_imp_nat_seg_image_inj_on [OF \<open>finite A\<close>]
   140   obtain f and n::nat where bij: "bij_betw f {i. i<n} A"
   148   obtain f and n::nat where bij: "bij_betw f {i. i<n} A"
   141     by (auto simp:bij_betw_def)
   149     by (auto simp: bij_betw_def)
   142   let ?f = "the_inv_into {i. i<n} f"
   150   let ?f = "the_inv_into {i. i<n} f"
   143   have "inj_on ?f A & ?f ` A = {i. i<n}"
   151   have "inj_on ?f A \<and> ?f ` A = {i. i<n}"
   144     by (fold bij_betw_def) (rule bij_betw_the_inv_into[OF bij])
   152     by (fold bij_betw_def) (rule bij_betw_the_inv_into[OF bij])
   145   thus ?thesis by blast
   153   then show ?thesis by blast
   146 qed
   154 qed
   147 
   155 
   148 lemma finite_Collect_less_nat [iff]:
   156 lemma finite_Collect_less_nat [iff]: "finite {n::nat. n < k}"
   149   "finite {n::nat. n < k}"
       
   150   by (fastforce simp: finite_conv_nat_seg_image)
   157   by (fastforce simp: finite_conv_nat_seg_image)
   151 
   158 
   152 lemma finite_Collect_le_nat [iff]:
   159 lemma finite_Collect_le_nat [iff]: "finite {n::nat. n \<le> k}"
   153   "finite {n::nat. n \<le> k}"
       
   154   by (simp add: le_eq_less_or_eq Collect_disj_eq)
   160   by (simp add: le_eq_less_or_eq Collect_disj_eq)
   155 
   161 
   156 
   162 
   157 subsubsection \<open>Finiteness and common set operations\<close>
   163 subsubsection \<open>Finiteness and common set operations\<close>
   158 
   164 
   159 lemma rev_finite_subset:
   165 lemma rev_finite_subset: "finite B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> finite A"
   160   "finite B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> finite A"
       
   161 proof (induct arbitrary: A rule: finite_induct)
   166 proof (induct arbitrary: A rule: finite_induct)
   162   case empty
   167   case empty
   163   then show ?case by simp
   168   then show ?case by simp
   164 next
   169 next
   165   case (insert x F A)
   170   case (insert x F A)
   166   have A: "A \<subseteq> insert x F" and r: "A - {x} \<subseteq> F \<Longrightarrow> finite (A - {x})" by fact+
   171   have A: "A \<subseteq> insert x F" and r: "A - {x} \<subseteq> F \<Longrightarrow> finite (A - {x})"
       
   172     by fact+
   167   show "finite A"
   173   show "finite A"
   168   proof cases
   174   proof cases
   169     assume x: "x \<in> A"
   175     assume x: "x \<in> A"
   170     with A have "A - {x} \<subseteq> F" by (simp add: subset_insert_iff)
   176     with A have "A - {x} \<subseteq> F" by (simp add: subset_insert_iff)
   171     with r have "finite (A - {x})" .
   177     with r have "finite (A - {x})" .
   172     hence "finite (insert x (A - {x}))" ..
   178     then have "finite (insert x (A - {x}))" ..
   173     also have "insert x (A - {x}) = A" using x by (rule insert_Diff)
   179     also have "insert x (A - {x}) = A"
       
   180       using x by (rule insert_Diff)
   174     finally show ?thesis .
   181     finally show ?thesis .
   175   next
   182   next
   176     show ?thesis when "A \<subseteq> F"
   183     show ?thesis when "A \<subseteq> F"
   177       using that by fact
   184       using that by fact
   178     assume "x \<notin> A"
   185     assume "x \<notin> A"
   179     with A show "A \<subseteq> F" by (simp add: subset_insert_iff)
   186     with A show "A \<subseteq> F"
   180   qed
   187       by (simp add: subset_insert_iff)
   181 qed
   188   qed
   182 
   189 qed
   183 lemma finite_subset:
   190 
   184   "A \<subseteq> B \<Longrightarrow> finite B \<Longrightarrow> finite A"
   191 lemma finite_subset: "A \<subseteq> B \<Longrightarrow> finite B \<Longrightarrow> finite A"
   185   by (rule rev_finite_subset)
   192   by (rule rev_finite_subset)
   186 
   193 
   187 lemma finite_UnI:
   194 lemma finite_UnI:
   188   assumes "finite F" and "finite G"
   195   assumes "finite F" and "finite G"
   189   shows "finite (F \<union> G)"
   196   shows "finite (F \<union> G)"
   190   using assms by induct simp_all
   197   using assms by induct simp_all
   191 
   198 
   192 lemma finite_Un [iff]:
   199 lemma finite_Un [iff]: "finite (F \<union> G) \<longleftrightarrow> finite F \<and> finite G"
   193   "finite (F \<union> G) \<longleftrightarrow> finite F \<and> finite G"
       
   194   by (blast intro: finite_UnI finite_subset [of _ "F \<union> G"])
   200   by (blast intro: finite_UnI finite_subset [of _ "F \<union> G"])
   195 
   201 
   196 lemma finite_insert [simp]: "finite (insert a A) \<longleftrightarrow> finite A"
   202 lemma finite_insert [simp]: "finite (insert a A) \<longleftrightarrow> finite A"
   197 proof -
   203 proof -
   198   have "finite {a} \<and> finite A \<longleftrightarrow> finite A" by simp
   204   have "finite {a} \<and> finite A \<longleftrightarrow> finite A" by simp
   199   then have "finite ({a} \<union> A) \<longleftrightarrow> finite A" by (simp only: finite_Un)
   205   then have "finite ({a} \<union> A) \<longleftrightarrow> finite A" by (simp only: finite_Un)
   200   then show ?thesis by simp
   206   then show ?thesis by simp
   201 qed
   207 qed
   202 
   208 
   203 lemma finite_Int [simp, intro]:
   209 lemma finite_Int [simp, intro]: "finite F \<or> finite G \<Longrightarrow> finite (F \<inter> G)"
   204   "finite F \<or> finite G \<Longrightarrow> finite (F \<inter> G)"
       
   205   by (blast intro: finite_subset)
   210   by (blast intro: finite_subset)
   206 
   211 
   207 lemma finite_Collect_conjI [simp, intro]:
   212 lemma finite_Collect_conjI [simp, intro]:
   208   "finite {x. P x} \<or> finite {x. Q x} \<Longrightarrow> finite {x. P x \<and> Q x}"
   213   "finite {x. P x} \<or> finite {x. Q x} \<Longrightarrow> finite {x. P x \<and> Q x}"
   209   by (simp add: Collect_conj_eq)
   214   by (simp add: Collect_conj_eq)
   210 
   215 
   211 lemma finite_Collect_disjI [simp]:
   216 lemma finite_Collect_disjI [simp]:
   212   "finite {x. P x \<or> Q x} \<longleftrightarrow> finite {x. P x} \<and> finite {x. Q x}"
   217   "finite {x. P x \<or> Q x} \<longleftrightarrow> finite {x. P x} \<and> finite {x. Q x}"
   213   by (simp add: Collect_disj_eq)
   218   by (simp add: Collect_disj_eq)
   214 
   219 
   215 lemma finite_Diff [simp, intro]:
   220 lemma finite_Diff [simp, intro]: "finite A \<Longrightarrow> finite (A - B)"
   216   "finite A \<Longrightarrow> finite (A - B)"
       
   217   by (rule finite_subset, rule Diff_subset)
   221   by (rule finite_subset, rule Diff_subset)
   218 
   222 
   219 lemma finite_Diff2 [simp]:
   223 lemma finite_Diff2 [simp]:
   220   assumes "finite B"
   224   assumes "finite B"
   221   shows "finite (A - B) \<longleftrightarrow> finite A"
   225   shows "finite (A - B) \<longleftrightarrow> finite A"
   222 proof -
   226 proof -
   223   have "finite A \<longleftrightarrow> finite((A - B) \<union> (A \<inter> B))" by (simp add: Un_Diff_Int)
   227   have "finite A \<longleftrightarrow> finite ((A - B) \<union> (A \<inter> B))"
   224   also have "\<dots> \<longleftrightarrow> finite (A - B)" using \<open>finite B\<close> by simp
   228     by (simp add: Un_Diff_Int)
       
   229   also have "\<dots> \<longleftrightarrow> finite (A - B)"
       
   230     using \<open>finite B\<close> by simp
   225   finally show ?thesis ..
   231   finally show ?thesis ..
   226 qed
   232 qed
   227 
   233 
   228 lemma finite_Diff_insert [iff]:
   234 lemma finite_Diff_insert [iff]: "finite (A - insert a B) \<longleftrightarrow> finite (A - B)"
   229   "finite (A - insert a B) \<longleftrightarrow> finite (A - B)"
       
   230 proof -
   235 proof -
   231   have "finite (A - B) \<longleftrightarrow> finite (A - B - {a})" by simp
   236   have "finite (A - B) \<longleftrightarrow> finite (A - B - {a})" by simp
   232   moreover have "A - insert a B = A - B - {a}" by auto
   237   moreover have "A - insert a B = A - B - {a}" by auto
   233   ultimately show ?thesis by simp
   238   ultimately show ?thesis by simp
   234 qed
   239 qed
   235 
   240 
   236 lemma finite_compl[simp]:
   241 lemma finite_compl [simp]:
   237   "finite (A :: 'a set) \<Longrightarrow> finite (- A) \<longleftrightarrow> finite (UNIV :: 'a set)"
   242   "finite (A :: 'a set) \<Longrightarrow> finite (- A) \<longleftrightarrow> finite (UNIV :: 'a set)"
   238   by (simp add: Compl_eq_Diff_UNIV)
   243   by (simp add: Compl_eq_Diff_UNIV)
   239 
   244 
   240 lemma finite_Collect_not[simp]:
   245 lemma finite_Collect_not [simp]:
   241   "finite {x :: 'a. P x} \<Longrightarrow> finite {x. \<not> P x} \<longleftrightarrow> finite (UNIV :: 'a set)"
   246   "finite {x :: 'a. P x} \<Longrightarrow> finite {x. \<not> P x} \<longleftrightarrow> finite (UNIV :: 'a set)"
   242   by (simp add: Collect_neg_eq)
   247   by (simp add: Collect_neg_eq)
   243 
   248 
   244 lemma finite_Union [simp, intro]:
   249 lemma finite_Union [simp, intro]:
   245   "finite A \<Longrightarrow> (\<And>M. M \<in> A \<Longrightarrow> finite M) \<Longrightarrow> finite(\<Union>A)"
   250   "finite A \<Longrightarrow> (\<And>M. M \<in> A \<Longrightarrow> finite M) \<Longrightarrow> finite (\<Union>A)"
   246   by (induct rule: finite_induct) simp_all
   251   by (induct rule: finite_induct) simp_all
   247 
   252 
   248 lemma finite_UN_I [intro]:
   253 lemma finite_UN_I [intro]:
   249   "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> finite (B a)) \<Longrightarrow> finite (\<Union>a\<in>A. B a)"
   254   "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> finite (B a)) \<Longrightarrow> finite (\<Union>a\<in>A. B a)"
   250   by (induct rule: finite_induct) simp_all
   255   by (induct rule: finite_induct) simp_all
   251 
   256 
   252 lemma finite_UN [simp]:
   257 lemma finite_UN [simp]: "finite A \<Longrightarrow> finite (UNION A B) \<longleftrightarrow> (\<forall>x\<in>A. finite (B x))"
   253   "finite A \<Longrightarrow> finite (UNION A B) \<longleftrightarrow> (\<forall>x\<in>A. finite (B x))"
       
   254   by (blast intro: finite_subset)
   258   by (blast intro: finite_subset)
   255 
   259 
   256 lemma finite_Inter [intro]:
   260 lemma finite_Inter [intro]: "\<exists>A\<in>M. finite A \<Longrightarrow> finite (\<Inter>M)"
   257   "\<exists>A\<in>M. finite A \<Longrightarrow> finite (\<Inter>M)"
       
   258   by (blast intro: Inter_lower finite_subset)
   261   by (blast intro: Inter_lower finite_subset)
   259 
   262 
   260 lemma finite_INT [intro]:
   263 lemma finite_INT [intro]: "\<exists>x\<in>I. finite (A x) \<Longrightarrow> finite (\<Inter>x\<in>I. A x)"
   261   "\<exists>x\<in>I. finite (A x) \<Longrightarrow> finite (\<Inter>x\<in>I. A x)"
       
   262   by (blast intro: INT_lower finite_subset)
   264   by (blast intro: INT_lower finite_subset)
   263 
   265 
   264 lemma finite_imageI [simp, intro]:
   266 lemma finite_imageI [simp, intro]: "finite F \<Longrightarrow> finite (h ` F)"
   265   "finite F \<Longrightarrow> finite (h ` F)"
       
   266   by (induct rule: finite_induct) simp_all
   267   by (induct rule: finite_induct) simp_all
   267 
   268 
   268 lemma finite_image_set [simp]:
   269 lemma finite_image_set [simp]: "finite {x. P x} \<Longrightarrow> finite {f x |x. P x}"
   269   "finite {x. P x} \<Longrightarrow> finite { f x | x. P x }"
       
   270   by (simp add: image_Collect [symmetric])
   270   by (simp add: image_Collect [symmetric])
   271 
   271 
   272 lemma finite_image_set2:
   272 lemma finite_image_set2:
   273   "finite {x. P x} \<Longrightarrow> finite {y. Q y} \<Longrightarrow> finite {f x y | x y. P x \<and> Q y}"
   273   "finite {x. P x} \<Longrightarrow> finite {y. Q y} \<Longrightarrow> finite {f x y |x y. P x \<and> Q y}"
   274   by (rule finite_subset [where B = "\<Union>x \<in> {x. P x}. \<Union>y \<in> {y. Q y}. {f x y}"]) auto
   274   by (rule finite_subset [where B = "\<Union>x \<in> {x. P x}. \<Union>y \<in> {y. Q y}. {f x y}"]) auto
   275 
   275 
   276 lemma finite_imageD:
   276 lemma finite_imageD:
   277   assumes "finite (f ` A)" and "inj_on f A"
   277   assumes "finite (f ` A)" and "inj_on f A"
   278   shows "finite A"
   278   shows "finite A"
   279 using assms
   279   using assms
   280 proof (induct "f ` A" arbitrary: A)
   280 proof (induct "f ` A" arbitrary: A)
   281   case empty then show ?case by simp
   281   case empty
       
   282   then show ?case by simp
   282 next
   283 next
   283   case (insert x B)
   284   case (insert x B)
   284   then have B_A: "insert x B = f ` A" by simp
   285   then have B_A: "insert x B = f ` A"
   285   then obtain y where "x = f y" and "y \<in> A" by blast
   286     by simp
   286   from B_A \<open>x \<notin> B\<close> have "B = f ` A - {x}" by blast
   287   then obtain y where "x = f y" and "y \<in> A"
   287   with B_A \<open>x \<notin> B\<close> \<open>x = f y\<close> \<open>inj_on f A\<close> \<open>y \<in> A\<close> have "B = f ` (A - {y})" 
   288     by blast
       
   289   from B_A \<open>x \<notin> B\<close> have "B = f ` A - {x}"
       
   290     by blast
       
   291   with B_A \<open>x \<notin> B\<close> \<open>x = f y\<close> \<open>inj_on f A\<close> \<open>y \<in> A\<close> have "B = f ` (A - {y})"
   288     by (simp add: inj_on_image_set_diff Set.Diff_subset)
   292     by (simp add: inj_on_image_set_diff Set.Diff_subset)
   289   moreover from \<open>inj_on f A\<close> have "inj_on f (A - {y})" by (rule inj_on_diff)
   293   moreover from \<open>inj_on f A\<close> have "inj_on f (A - {y})"
   290   ultimately have "finite (A - {y})" by (rule insert.hyps)
   294     by (rule inj_on_diff)
   291   then show "finite A" by simp
   295   ultimately have "finite (A - {y})"
   292 qed
   296     by (rule insert.hyps)
   293 
   297   then show "finite A"
   294 lemma finite_image_iff:
   298     by simp
   295   assumes "inj_on f A"
   299 qed
   296   shows "finite (f ` A) \<longleftrightarrow> finite A"
   300 
   297 using assms finite_imageD by blast
   301 lemma finite_image_iff: "inj_on f A \<Longrightarrow> finite (f ` A) \<longleftrightarrow> finite A"
   298 
   302   using finite_imageD by blast
   299 lemma finite_surj:
   303 
   300   "finite A \<Longrightarrow> B \<subseteq> f ` A \<Longrightarrow> finite B"
   304 lemma finite_surj: "finite A \<Longrightarrow> B \<subseteq> f ` A \<Longrightarrow> finite B"
   301   by (erule finite_subset) (rule finite_imageI)
   305   by (erule finite_subset) (rule finite_imageI)
   302 
   306 
   303 lemma finite_range_imageI:
   307 lemma finite_range_imageI: "finite (range g) \<Longrightarrow> finite (range (\<lambda>x. f (g x)))"
   304   "finite (range g) \<Longrightarrow> finite (range (\<lambda>x. f (g x)))"
       
   305   by (drule finite_imageI) (simp add: range_composition)
   308   by (drule finite_imageI) (simp add: range_composition)
   306 
   309 
   307 lemma finite_subset_image:
   310 lemma finite_subset_image:
   308   assumes "finite B"
   311   assumes "finite B"
   309   shows "B \<subseteq> f ` A \<Longrightarrow> \<exists>C\<subseteq>A. finite C \<and> B = f ` C"
   312   shows "B \<subseteq> f ` A \<Longrightarrow> \<exists>C\<subseteq>A. finite C \<and> B = f ` C"
   310 using assms
   313   using assms
   311 proof induct
   314 proof induct
   312   case empty then show ?case by simp
   315   case empty
   313 next
   316   then show ?case by simp
   314   case insert then show ?case
   317 next
   315     by (clarsimp simp del: image_insert simp add: image_insert [symmetric])
   318   case insert
   316        blast
   319   then show ?case
   317 qed
   320     by (clarsimp simp del: image_insert simp add: image_insert [symmetric]) blast
   318 
   321 qed
   319 lemma finite_vimage_IntI:
   322 
   320   "finite F \<Longrightarrow> inj_on h A \<Longrightarrow> finite (h -` F \<inter> A)"
   323 lemma finite_vimage_IntI: "finite F \<Longrightarrow> inj_on h A \<Longrightarrow> finite (h -` F \<inter> A)"
   321   apply (induct rule: finite_induct)
   324   apply (induct rule: finite_induct)
   322    apply simp_all
   325    apply simp_all
   323   apply (subst vimage_insert)
   326   apply (subst vimage_insert)
   324   apply (simp add: finite_subset [OF inj_on_vimage_singleton] Int_Un_distrib2)
   327   apply (simp add: finite_subset [OF inj_on_vimage_singleton] Int_Un_distrib2)
   325   done
   328   done
   332     by blast
   335     by blast
   333   show ?thesis
   336   show ?thesis
   334     by (simp only: * assms finite_UN_I)
   337     by (simp only: * assms finite_UN_I)
   335 qed
   338 qed
   336 
   339 
   337 lemma finite_vimageI:
   340 lemma finite_vimageI: "finite F \<Longrightarrow> inj h \<Longrightarrow> finite (h -` F)"
   338   "finite F \<Longrightarrow> inj h \<Longrightarrow> finite (h -` F)"
       
   339   using finite_vimage_IntI[of F h UNIV] by auto
   341   using finite_vimage_IntI[of F h UNIV] by auto
   340 
   342 
   341 lemma finite_vimageD': "\<lbrakk> finite (f -` A); A \<subseteq> range f \<rbrakk> \<Longrightarrow> finite A"
   343 lemma finite_vimageD': "finite (f -` A) \<Longrightarrow> A \<subseteq> range f \<Longrightarrow> finite A"
   342 by(auto simp add: subset_image_iff intro: finite_subset[rotated])
   344   by (auto simp add: subset_image_iff intro: finite_subset[rotated])
   343 
   345 
   344 lemma finite_vimageD: "\<lbrakk> finite (h -` F); surj h \<rbrakk> \<Longrightarrow> finite F"
   346 lemma finite_vimageD: "finite (h -` F) \<Longrightarrow> surj h \<Longrightarrow> finite F"
   345 by(auto dest: finite_vimageD')
   347   by (auto dest: finite_vimageD')
   346 
   348 
   347 lemma finite_vimage_iff: "bij h \<Longrightarrow> finite (h -` F) \<longleftrightarrow> finite F"
   349 lemma finite_vimage_iff: "bij h \<Longrightarrow> finite (h -` F) \<longleftrightarrow> finite F"
   348   unfolding bij_def by (auto elim: finite_vimageD finite_vimageI)
   350   unfolding bij_def by (auto elim: finite_vimageD finite_vimageI)
   349 
   351 
   350 lemma finite_Collect_bex [simp]:
   352 lemma finite_Collect_bex [simp]:
   357 
   359 
   358 lemma finite_Collect_bounded_ex [simp]:
   360 lemma finite_Collect_bounded_ex [simp]:
   359   assumes "finite {y. P y}"
   361   assumes "finite {y. P y}"
   360   shows "finite {x. \<exists>y. P y \<and> Q x y} \<longleftrightarrow> (\<forall>y. P y \<longrightarrow> finite {x. Q x y})"
   362   shows "finite {x. \<exists>y. P y \<and> Q x y} \<longleftrightarrow> (\<forall>y. P y \<longrightarrow> finite {x. Q x y})"
   361 proof -
   363 proof -
   362   have "{x. EX y. P y & Q x y} = (\<Union>y\<in>{y. P y}. {x. Q x y})" by auto
   364   have "{x. \<exists>y. P y \<and> Q x y} = (\<Union>y\<in>{y. P y}. {x. Q x y})"
   363   with assms show ?thesis by simp
   365     by auto
   364 qed
   366   with assms show ?thesis
   365 
   367     by simp
   366 lemma finite_Plus:
   368 qed
   367   "finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A <+> B)"
   369 
       
   370 lemma finite_Plus: "finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A <+> B)"
   368   by (simp add: Plus_def)
   371   by (simp add: Plus_def)
   369 
   372 
   370 lemma finite_PlusD: 
   373 lemma finite_PlusD:
   371   fixes A :: "'a set" and B :: "'b set"
   374   fixes A :: "'a set" and B :: "'b set"
   372   assumes fin: "finite (A <+> B)"
   375   assumes fin: "finite (A <+> B)"
   373   shows "finite A" "finite B"
   376   shows "finite A" "finite B"
   374 proof -
   377 proof -
   375   have "Inl ` A \<subseteq> A <+> B" by auto
   378   have "Inl ` A \<subseteq> A <+> B"
   376   then have "finite (Inl ` A :: ('a + 'b) set)" using fin by (rule finite_subset)
   379     by auto
   377   then show "finite A" by (rule finite_imageD) (auto intro: inj_onI)
   380   then have "finite (Inl ` A :: ('a + 'b) set)"
   378 next
   381     using fin by (rule finite_subset)
   379   have "Inr ` B \<subseteq> A <+> B" by auto
   382   then show "finite A"
   380   then have "finite (Inr ` B :: ('a + 'b) set)" using fin by (rule finite_subset)
   383     by (rule finite_imageD) (auto intro: inj_onI)
   381   then show "finite B" by (rule finite_imageD) (auto intro: inj_onI)
   384 next
   382 qed
   385   have "Inr ` B \<subseteq> A <+> B"
   383 
   386     by auto
   384 lemma finite_Plus_iff [simp]:
   387   then have "finite (Inr ` B :: ('a + 'b) set)"
   385   "finite (A <+> B) \<longleftrightarrow> finite A \<and> finite B"
   388     using fin by (rule finite_subset)
       
   389   then show "finite B"
       
   390     by (rule finite_imageD) (auto intro: inj_onI)
       
   391 qed
       
   392 
       
   393 lemma finite_Plus_iff [simp]: "finite (A <+> B) \<longleftrightarrow> finite A \<and> finite B"
   386   by (auto intro: finite_PlusD finite_Plus)
   394   by (auto intro: finite_PlusD finite_Plus)
   387 
   395 
   388 lemma finite_Plus_UNIV_iff [simp]:
   396 lemma finite_Plus_UNIV_iff [simp]:
   389   "finite (UNIV :: ('a + 'b) set) \<longleftrightarrow> finite (UNIV :: 'a set) \<and> finite (UNIV :: 'b set)"
   397   "finite (UNIV :: ('a + 'b) set) \<longleftrightarrow> finite (UNIV :: 'a set) \<and> finite (UNIV :: 'b set)"
   390   by (subst UNIV_Plus_UNIV [symmetric]) (rule finite_Plus_iff)
   398   by (subst UNIV_Plus_UNIV [symmetric]) (rule finite_Plus_iff)
   391 
   399 
   392 lemma finite_SigmaI [simp, intro]:
   400 lemma finite_SigmaI [simp, intro]:
   393   "finite A \<Longrightarrow> (\<And>a. a\<in>A \<Longrightarrow> finite (B a)) ==> finite (SIGMA a:A. B a)"
   401   "finite A \<Longrightarrow> (\<And>a. a\<in>A \<Longrightarrow> finite (B a)) \<Longrightarrow> finite (SIGMA a:A. B a)"
   394   by (unfold Sigma_def) blast
   402   unfolding Sigma_def by blast
   395 
   403 
   396 lemma finite_SigmaI2:
   404 lemma finite_SigmaI2:
   397   assumes "finite {x\<in>A. B x \<noteq> {}}"
   405   assumes "finite {x\<in>A. B x \<noteq> {}}"
   398   and "\<And>a. a \<in> A \<Longrightarrow> finite (B a)"
   406   and "\<And>a. a \<in> A \<Longrightarrow> finite (B a)"
   399   shows "finite (Sigma A B)"
   407   shows "finite (Sigma A B)"
   400 proof -
   408 proof -
   401   from assms have "finite (Sigma {x\<in>A. B x \<noteq> {}} B)" by auto
   409   from assms have "finite (Sigma {x\<in>A. B x \<noteq> {}} B)"
   402   also have "Sigma {x:A. B x \<noteq> {}} B = Sigma A B" by auto
   410     by auto
       
   411   also have "Sigma {x:A. B x \<noteq> {}} B = Sigma A B"
       
   412     by auto
   403   finally show ?thesis .
   413   finally show ?thesis .
   404 qed
   414 qed
   405 
   415 
   406 lemma finite_cartesian_product:
   416 lemma finite_cartesian_product: "finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<times> B)"
   407   "finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<times> B)"
       
   408   by (rule finite_SigmaI)
   417   by (rule finite_SigmaI)
   409 
   418 
   410 lemma finite_Prod_UNIV:
   419 lemma finite_Prod_UNIV:
   411   "finite (UNIV :: 'a set) \<Longrightarrow> finite (UNIV :: 'b set) \<Longrightarrow> finite (UNIV :: ('a \<times> 'b) set)"
   420   "finite (UNIV :: 'a set) \<Longrightarrow> finite (UNIV :: 'b set) \<Longrightarrow> finite (UNIV :: ('a \<times> 'b) set)"
   412   by (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product)
   421   by (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product)
   415   assumes "finite (A \<times> B)" and "B \<noteq> {}"
   424   assumes "finite (A \<times> B)" and "B \<noteq> {}"
   416   shows "finite A"
   425   shows "finite A"
   417 proof -
   426 proof -
   418   from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
   427   from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
   419     by (auto simp add: finite_conv_nat_seg_image)
   428     by (auto simp add: finite_conv_nat_seg_image)
   420   then have "fst ` (A \<times> B) = fst ` f ` {i::nat. i < n}" by simp
   429   then have "fst ` (A \<times> B) = fst ` f ` {i::nat. i < n}"
       
   430     by simp
   421   with \<open>B \<noteq> {}\<close> have "A = (fst \<circ> f) ` {i::nat. i < n}"
   431   with \<open>B \<noteq> {}\<close> have "A = (fst \<circ> f) ` {i::nat. i < n}"
   422     by (simp add: image_comp)
   432     by (simp add: image_comp)
   423   then have "\<exists>n f. A = f ` {i::nat. i < n}" by blast
   433   then have "\<exists>n f. A = f ` {i::nat. i < n}"
       
   434     by blast
   424   then show ?thesis
   435   then show ?thesis
   425     by (auto simp add: finite_conv_nat_seg_image)
   436     by (auto simp add: finite_conv_nat_seg_image)
   426 qed
   437 qed
   427 
   438 
   428 lemma finite_cartesian_productD2:
   439 lemma finite_cartesian_productD2:
   429   assumes "finite (A \<times> B)" and "A \<noteq> {}"
   440   assumes "finite (A \<times> B)" and "A \<noteq> {}"
   430   shows "finite B"
   441   shows "finite B"
   431 proof -
   442 proof -
   432   from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
   443   from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
   433     by (auto simp add: finite_conv_nat_seg_image)
   444     by (auto simp add: finite_conv_nat_seg_image)
   434   then have "snd ` (A \<times> B) = snd ` f ` {i::nat. i < n}" by simp
   445   then have "snd ` (A \<times> B) = snd ` f ` {i::nat. i < n}"
       
   446     by simp
   435   with \<open>A \<noteq> {}\<close> have "B = (snd \<circ> f) ` {i::nat. i < n}"
   447   with \<open>A \<noteq> {}\<close> have "B = (snd \<circ> f) ` {i::nat. i < n}"
   436     by (simp add: image_comp)
   448     by (simp add: image_comp)
   437   then have "\<exists>n f. B = f ` {i::nat. i < n}" by blast
   449   then have "\<exists>n f. B = f ` {i::nat. i < n}"
       
   450     by blast
   438   then show ?thesis
   451   then show ?thesis
   439     by (auto simp add: finite_conv_nat_seg_image)
   452     by (auto simp add: finite_conv_nat_seg_image)
   440 qed
   453 qed
   441 
   454 
   442 lemma finite_cartesian_product_iff:
   455 lemma finite_cartesian_product_iff:
   443   "finite (A \<times> B) \<longleftrightarrow> (A = {} \<or> B = {} \<or> (finite A \<and> finite B))"
   456   "finite (A \<times> B) \<longleftrightarrow> (A = {} \<or> B = {} \<or> (finite A \<and> finite B))"
   444   by (auto dest: finite_cartesian_productD1 finite_cartesian_productD2 finite_cartesian_product)
   457   by (auto dest: finite_cartesian_productD1 finite_cartesian_productD2 finite_cartesian_product)
   445 
   458 
   446 lemma finite_prod: 
   459 lemma finite_prod:
   447   "finite (UNIV :: ('a \<times> 'b) set) \<longleftrightarrow> finite (UNIV :: 'a set) \<and> finite (UNIV :: 'b set)"
   460   "finite (UNIV :: ('a \<times> 'b) set) \<longleftrightarrow> finite (UNIV :: 'a set) \<and> finite (UNIV :: 'b set)"
   448   using finite_cartesian_product_iff[of UNIV UNIV] by simp
   461   using finite_cartesian_product_iff[of UNIV UNIV] by simp
   449 
   462 
   450 lemma finite_Pow_iff [iff]:
   463 lemma finite_Pow_iff [iff]: "finite (Pow A) \<longleftrightarrow> finite A"
   451   "finite (Pow A) \<longleftrightarrow> finite A"
       
   452 proof
   464 proof
   453   assume "finite (Pow A)"
   465   assume "finite (Pow A)"
   454   then have "finite ((%x. {x}) ` A)" by (blast intro: finite_subset)
   466   then have "finite ((\<lambda>x. {x}) ` A)"
   455   then show "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp
   467     by (blast intro: finite_subset)
       
   468   then show "finite A"
       
   469     by (rule finite_imageD [unfolded inj_on_def]) simp
   456 next
   470 next
   457   assume "finite A"
   471   assume "finite A"
   458   then show "finite (Pow A)"
   472   then show "finite (Pow A)"
   459     by induct (simp_all add: Pow_insert)
   473     by induct (simp_all add: Pow_insert)
   460 qed
   474 qed
   461 
   475 
   462 corollary finite_Collect_subsets [simp, intro]:
   476 corollary finite_Collect_subsets [simp, intro]: "finite A \<Longrightarrow> finite {B. B \<subseteq> A}"
   463   "finite A \<Longrightarrow> finite {B. B \<subseteq> A}"
       
   464   by (simp add: Pow_def [symmetric])
   477   by (simp add: Pow_def [symmetric])
   465 
   478 
   466 lemma finite_set: "finite (UNIV :: 'a set set) \<longleftrightarrow> finite (UNIV :: 'a set)"
   479 lemma finite_set: "finite (UNIV :: 'a set set) \<longleftrightarrow> finite (UNIV :: 'a set)"
   467 by(simp only: finite_Pow_iff Pow_UNIV[symmetric])
   480   by (simp only: finite_Pow_iff Pow_UNIV[symmetric])
   468 
   481 
   469 lemma finite_UnionD: "finite(\<Union>A) \<Longrightarrow> finite A"
   482 lemma finite_UnionD: "finite (\<Union>A) \<Longrightarrow> finite A"
   470   by (blast intro: finite_subset [OF subset_Pow_Union])
   483   by (blast intro: finite_subset [OF subset_Pow_Union])
   471 
   484 
   472 lemma finite_set_of_finite_funs: assumes "finite A" "finite B"
   485 lemma finite_set_of_finite_funs:
   473 shows "finite{f. \<forall>x. (x \<in> A \<longrightarrow> f x \<in> B) \<and> (x \<notin> A \<longrightarrow> f x = d)}" (is "finite ?S")
   486   assumes "finite A" "finite B"
   474 proof-
   487   shows "finite {f. \<forall>x. (x \<in> A \<longrightarrow> f x \<in> B) \<and> (x \<notin> A \<longrightarrow> f x = d)}" (is "finite ?S")
       
   488 proof -
   475   let ?F = "\<lambda>f. {(a,b). a \<in> A \<and> b = f a}"
   489   let ?F = "\<lambda>f. {(a,b). a \<in> A \<and> b = f a}"
   476   have "?F ` ?S \<subseteq> Pow(A \<times> B)" by auto
   490   have "?F ` ?S \<subseteq> Pow(A \<times> B)"
   477   from finite_subset[OF this] assms have 1: "finite (?F ` ?S)" by simp
   491     by auto
       
   492   from finite_subset[OF this] assms have 1: "finite (?F ` ?S)"
       
   493     by simp
   478   have 2: "inj_on ?F ?S"
   494   have 2: "inj_on ?F ?S"
   479     by(fastforce simp add: inj_on_def set_eq_iff fun_eq_iff)
   495     by (fastforce simp add: inj_on_def set_eq_iff fun_eq_iff)
   480   show ?thesis by(rule finite_imageD[OF 1 2])
   496   show ?thesis
       
   497     by (rule finite_imageD [OF 1 2])
   481 qed
   498 qed
   482 
   499 
   483 lemma not_finite_existsD:
   500 lemma not_finite_existsD:
   484   assumes "\<not> finite {a. P a}"
   501   assumes "\<not> finite {a. P a}"
   485   shows "\<exists>a. P a"
   502   shows "\<exists>a. P a"
   486 proof (rule classical)
   503 proof (rule classical)
   487   assume "\<not> (\<exists>a. P a)"
   504   assume "\<not> ?thesis"
   488   with assms show ?thesis by auto
   505   with assms show ?thesis by auto
   489 qed
   506 qed
   490 
   507 
   491 
   508 
   492 subsubsection \<open>Further induction rules on finite sets\<close>
   509 subsubsection \<open>Further induction rules on finite sets\<close>
   494 lemma finite_ne_induct [case_names singleton insert, consumes 2]:
   511 lemma finite_ne_induct [case_names singleton insert, consumes 2]:
   495   assumes "finite F" and "F \<noteq> {}"
   512   assumes "finite F" and "F \<noteq> {}"
   496   assumes "\<And>x. P {x}"
   513   assumes "\<And>x. P {x}"
   497     and "\<And>x F. finite F \<Longrightarrow> F \<noteq> {} \<Longrightarrow> x \<notin> F \<Longrightarrow> P F  \<Longrightarrow> P (insert x F)"
   514     and "\<And>x F. finite F \<Longrightarrow> F \<noteq> {} \<Longrightarrow> x \<notin> F \<Longrightarrow> P F  \<Longrightarrow> P (insert x F)"
   498   shows "P F"
   515   shows "P F"
   499 using assms
   516   using assms
   500 proof induct
   517 proof induct
   501   case empty then show ?case by simp
   518   case empty
   502 next
   519   then show ?case by simp
   503   case (insert x F) then show ?case by cases auto
   520 next
       
   521   case (insert x F)
       
   522   then show ?case by cases auto
   504 qed
   523 qed
   505 
   524 
   506 lemma finite_subset_induct [consumes 2, case_names empty insert]:
   525 lemma finite_subset_induct [consumes 2, case_names empty insert]:
   507   assumes "finite F" and "F \<subseteq> A"
   526   assumes "finite F" and "F \<subseteq> A"
   508   assumes empty: "P {}"
   527   assumes empty: "P {}"
   509     and insert: "\<And>a F. finite F \<Longrightarrow> a \<in> A \<Longrightarrow> a \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert a F)"
   528     and insert: "\<And>a F. finite F \<Longrightarrow> a \<in> A \<Longrightarrow> a \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert a F)"
   510   shows "P F"
   529   shows "P F"
   511 using \<open>finite F\<close> \<open>F \<subseteq> A\<close>
   530   using \<open>finite F\<close> \<open>F \<subseteq> A\<close>
   512 proof induct
   531 proof induct
   513   show "P {}" by fact
   532   show "P {}" by fact
   514 next
   533 next
   515   fix x F
   534   fix x F
   516   assume "finite F" and "x \<notin> F" and
   535   assume "finite F" and "x \<notin> F" and P: "F \<subseteq> A \<Longrightarrow> P F" and i: "insert x F \<subseteq> A"
   517     P: "F \<subseteq> A \<Longrightarrow> P F" and i: "insert x F \<subseteq> A"
       
   518   show "P (insert x F)"
   536   show "P (insert x F)"
   519   proof (rule insert)
   537   proof (rule insert)
   520     from i show "x \<in> A" by blast
   538     from i show "x \<in> A" by blast
   521     from i have "F \<subseteq> A" by blast
   539     from i have "F \<subseteq> A" by blast
   522     with P show "P F" .
   540     with P show "P F" .
   529   assumes "finite A"
   547   assumes "finite A"
   530   assumes "P A"
   548   assumes "P A"
   531     and remove: "\<And>a A. finite A \<Longrightarrow> a \<in> A \<Longrightarrow> P A \<Longrightarrow> P (A - {a})"
   549     and remove: "\<And>a A. finite A \<Longrightarrow> a \<in> A \<Longrightarrow> P A \<Longrightarrow> P (A - {a})"
   532   shows "P {}"
   550   shows "P {}"
   533 proof -
   551 proof -
   534   have "\<And>B. B \<subseteq> A \<Longrightarrow> P (A - B)"
   552   have "P (A - B)" if "B \<subseteq> A" for B :: "'a set"
   535   proof -
   553   proof -
   536     fix B :: "'a set"
   554     from \<open>finite A\<close> that have "finite B"
   537     assume "B \<subseteq> A"
   555       by (rule rev_finite_subset)
   538     with \<open>finite A\<close> have "finite B" by (rule rev_finite_subset)
       
   539     from this \<open>B \<subseteq> A\<close> show "P (A - B)"
   556     from this \<open>B \<subseteq> A\<close> show "P (A - B)"
   540     proof induct
   557     proof induct
   541       case empty
   558       case empty
   542       from \<open>P A\<close> show ?case by simp
   559       from \<open>P A\<close> show ?case by simp
   543     next
   560     next
   544       case (insert b B)
   561       case (insert b B)
   545       have "P (A - B - {b})"
   562       have "P (A - B - {b})"
   546       proof (rule remove)
   563       proof (rule remove)
   547         from \<open>finite A\<close> show "finite (A - B)" by induct auto
   564         from \<open>finite A\<close> show "finite (A - B)"
   548         from insert show "b \<in> A - B" by simp
   565           by induct auto
   549         from insert show "P (A - B)" by simp
   566         from insert show "b \<in> A - B"
       
   567           by simp
       
   568         from insert show "P (A - B)"
       
   569           by simp
   550       qed
   570       qed
   551       also have "A - B - {b} = A - insert b B" by (rule Diff_insert [symmetric])
   571       also have "A - B - {b} = A - insert b B"
       
   572         by (rule Diff_insert [symmetric])
   552       finally show ?case .
   573       finally show ?case .
   553     qed
   574     qed
   554   qed
   575   qed
   555   then have "P (A - A)" by blast
   576   then have "P (A - A)" by blast
   556   then show ?thesis by simp
   577   then show ?thesis by simp
   557 qed
   578 qed
   558 
   579 
   559 lemma finite_update_induct [consumes 1, case_names const update]:
   580 lemma finite_update_induct [consumes 1, case_names const update]:
   560   assumes finite: "finite {a. f a \<noteq> c}"
   581   assumes finite: "finite {a. f a \<noteq> c}"
   561   assumes const: "P (\<lambda>a. c)"
   582     and const: "P (\<lambda>a. c)"
   562   assumes update: "\<And>a b f. finite {a. f a \<noteq> c} \<Longrightarrow> f a = c \<Longrightarrow> b \<noteq> c \<Longrightarrow> P f \<Longrightarrow> P (f(a := b))"
   583     and update: "\<And>a b f. finite {a. f a \<noteq> c} \<Longrightarrow> f a = c \<Longrightarrow> b \<noteq> c \<Longrightarrow> P f \<Longrightarrow> P (f(a := b))"
   563   shows "P f"
   584   shows "P f"
   564 using finite proof (induct "{a. f a \<noteq> c}" arbitrary: f)
   585   using finite
   565   case empty with const show ?case by simp
   586 proof (induct "{a. f a \<noteq> c}" arbitrary: f)
       
   587   case empty
       
   588   with const show ?case by simp
   566 next
   589 next
   567   case (insert a A)
   590   case (insert a A)
   568   then have "A = {a'. (f(a := c)) a' \<noteq> c}" and "f a \<noteq> c"
   591   then have "A = {a'. (f(a := c)) a' \<noteq> c}" and "f a \<noteq> c"
   569     by auto
   592     by auto
   570   with \<open>finite A\<close> have "finite {a'. (f(a := c)) a' \<noteq> c}"
   593   with \<open>finite A\<close> have "finite {a'. (f(a := c)) a' \<noteq> c}"
   571     by simp
   594     by simp
   572   have "(f(a := c)) a = c"
   595   have "(f(a := c)) a = c"
   573     by simp
   596     by simp
   574   from insert \<open>A = {a'. (f(a := c)) a' \<noteq> c}\<close> have "P (f(a := c))"
   597   from insert \<open>A = {a'. (f(a := c)) a' \<noteq> c}\<close> have "P (f(a := c))"
   575     by simp
   598     by simp
   576   with \<open>finite {a'. (f(a := c)) a' \<noteq> c}\<close> \<open>(f(a := c)) a = c\<close> \<open>f a \<noteq> c\<close> have "P ((f(a := c))(a := f a))"
   599   with \<open>finite {a'. (f(a := c)) a' \<noteq> c}\<close> \<open>(f(a := c)) a = c\<close> \<open>f a \<noteq> c\<close>
       
   600   have "P ((f(a := c))(a := f a))"
   577     by (rule update)
   601     by (rule update)
   578   then show ?case by simp
   602   then show ?case by simp
   579 qed
   603 qed
   580 
   604 
   581 
   605 
   582 subsection \<open>Class \<open>finite\<close>\<close>
   606 subsection \<open>Class \<open>finite\<close>\<close>
   583 
   607 
   584 class finite =
   608 class finite = assumes finite_UNIV: "finite (UNIV :: 'a set)"
   585   assumes finite_UNIV: "finite (UNIV :: 'a set)"
       
   586 begin
   609 begin
   587 
   610 
   588 lemma finite [simp]: "finite (A :: 'a set)"
   611 lemma finite [simp]: "finite (A :: 'a set)"
   589   by (rule subset_UNIV finite_UNIV finite_subset)+
   612   by (rule subset_UNIV finite_UNIV finite_subset)+
   590 
   613 
   594 end
   617 end
   595 
   618 
   596 instance prod :: (finite, finite) finite
   619 instance prod :: (finite, finite) finite
   597   by standard (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite)
   620   by standard (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite)
   598 
   621 
   599 lemma inj_graph: "inj (%f. {(x, y). y = f x})"
   622 lemma inj_graph: "inj (\<lambda>f. {(x, y). y = f x})"
   600   by (rule inj_onI, auto simp add: set_eq_iff fun_eq_iff)
   623   by (rule inj_onI) (auto simp add: set_eq_iff fun_eq_iff)
   601 
   624 
   602 instance "fun" :: (finite, finite) finite
   625 instance "fun" :: (finite, finite) finite
   603 proof
   626 proof
   604   show "finite (UNIV :: ('a => 'b) set)"
   627   show "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
   605   proof (rule finite_imageD)
   628   proof (rule finite_imageD)
   606     let ?graph = "%f::'a => 'b. {(x, y). y = f x}"
   629     let ?graph = "\<lambda>f::'a \<Rightarrow> 'b. {(x, y). y = f x}"
   607     have "range ?graph \<subseteq> Pow UNIV" by simp
   630     have "range ?graph \<subseteq> Pow UNIV"
       
   631       by simp
   608     moreover have "finite (Pow (UNIV :: ('a * 'b) set))"
   632     moreover have "finite (Pow (UNIV :: ('a * 'b) set))"
   609       by (simp only: finite_Pow_iff finite)
   633       by (simp only: finite_Pow_iff finite)
   610     ultimately show "finite (range ?graph)"
   634     ultimately show "finite (range ?graph)"
   611       by (rule finite_subset)
   635       by (rule finite_subset)
   612     show "inj ?graph" by (rule inj_graph)
   636     show "inj ?graph"
       
   637       by (rule inj_graph)
   613   qed
   638   qed
   614 qed
   639 qed
   615 
   640 
   616 instance bool :: finite
   641 instance bool :: finite
   617   by standard (simp add: UNIV_bool)
   642   by standard (simp add: UNIV_bool)
   627 
   652 
   628 
   653 
   629 subsection \<open>A basic fold functional for finite sets\<close>
   654 subsection \<open>A basic fold functional for finite sets\<close>
   630 
   655 
   631 text \<open>The intended behaviour is
   656 text \<open>The intended behaviour is
   632 \<open>fold f z {x\<^sub>1, ..., x\<^sub>n} = f x\<^sub>1 (\<dots> (f x\<^sub>n z)\<dots>)\<close>
   657   \<open>fold f z {x\<^sub>1, \<dots>, x\<^sub>n} = f x\<^sub>1 (\<dots> (f x\<^sub>n z)\<dots>)\<close>
   633 if \<open>f\<close> is ``left-commutative'':
   658   if \<open>f\<close> is ``left-commutative'':
   634 \<close>
   659 \<close>
   635 
   660 
   636 locale comp_fun_commute =
   661 locale comp_fun_commute =
   637   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
   662   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
   638   assumes comp_fun_commute: "f y \<circ> f x = f x \<circ> f y"
   663   assumes comp_fun_commute: "f y \<circ> f x = f x \<circ> f y"
   639 begin
   664 begin
   640 
   665 
   641 lemma fun_left_comm: "f y (f x z) = f x (f y z)"
   666 lemma fun_left_comm: "f y (f x z) = f x (f y z)"
   642   using comp_fun_commute by (simp add: fun_eq_iff)
   667   using comp_fun_commute by (simp add: fun_eq_iff)
   643 
   668 
   644 lemma commute_left_comp:
   669 lemma commute_left_comp: "f y \<circ> (f x \<circ> g) = f x \<circ> (f y \<circ> g)"
   645   "f y \<circ> (f x \<circ> g) = f x \<circ> (f y \<circ> g)"
       
   646   by (simp add: o_assoc comp_fun_commute)
   670   by (simp add: o_assoc comp_fun_commute)
   647 
   671 
   648 end
   672 end
   649 
   673 
   650 inductive fold_graph :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool"
   674 inductive fold_graph :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool"
   651 for f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" and z :: 'b where
   675   for f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" and z :: 'b
   652   emptyI [intro]: "fold_graph f z {} z" |
   676 where
   653   insertI [intro]: "x \<notin> A \<Longrightarrow> fold_graph f z A y
   677   emptyI [intro]: "fold_graph f z {} z"
   654       \<Longrightarrow> fold_graph f z (insert x A) (f x y)"
   678 | insertI [intro]: "x \<notin> A \<Longrightarrow> fold_graph f z A y \<Longrightarrow> fold_graph f z (insert x A) (f x y)"
   655 
   679 
   656 inductive_cases empty_fold_graphE [elim!]: "fold_graph f z {} x"
   680 inductive_cases empty_fold_graphE [elim!]: "fold_graph f z {} x"
   657 
   681 
   658 definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b" where
   682 definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b"
   659   "fold f z A = (if finite A then (THE y. fold_graph f z A y) else z)"
   683   where "fold f z A = (if finite A then (THE y. fold_graph f z A y) else z)"
   660 
   684 
   661 text\<open>A tempting alternative for the definiens is
   685 text \<open>
   662 @{term "if finite A then THE y. fold_graph f z A y else e"}.
   686   A tempting alternative for the definiens is
   663 It allows the removal of finiteness assumptions from the theorems
   687   @{term "if finite A then THE y. fold_graph f z A y else e"}.
   664 \<open>fold_comm\<close>, \<open>fold_reindex\<close> and \<open>fold_distrib\<close>.
   688   It allows the removal of finiteness assumptions from the theorems
   665 The proofs become ugly. It is not worth the effort. (???)\<close>
   689   \<open>fold_comm\<close>, \<open>fold_reindex\<close> and \<open>fold_distrib\<close>.
       
   690   The proofs become ugly. It is not worth the effort. (???)
       
   691 \<close>
   666 
   692 
   667 lemma finite_imp_fold_graph: "finite A \<Longrightarrow> \<exists>x. fold_graph f z A x"
   693 lemma finite_imp_fold_graph: "finite A \<Longrightarrow> \<exists>x. fold_graph f z A x"
   668 by (induct rule: finite_induct) auto
   694   by (induct rule: finite_induct) auto
   669 
   695 
   670 
   696 
   671 subsubsection\<open>From @{const fold_graph} to @{term fold}\<close>
   697 subsubsection \<open>From @{const fold_graph} to @{term fold}\<close>
   672 
   698 
   673 context comp_fun_commute
   699 context comp_fun_commute
   674 begin
   700 begin
   675 
   701 
   676 lemma fold_graph_finite:
   702 lemma fold_graph_finite:
   679   using assms by induct simp_all
   705   using assms by induct simp_all
   680 
   706 
   681 lemma fold_graph_insertE_aux:
   707 lemma fold_graph_insertE_aux:
   682   "fold_graph f z A y \<Longrightarrow> a \<in> A \<Longrightarrow> \<exists>y'. y = f a y' \<and> fold_graph f z (A - {a}) y'"
   708   "fold_graph f z A y \<Longrightarrow> a \<in> A \<Longrightarrow> \<exists>y'. y = f a y' \<and> fold_graph f z (A - {a}) y'"
   683 proof (induct set: fold_graph)
   709 proof (induct set: fold_graph)
   684   case (insertI x A y) show ?case
   710   case emptyI
       
   711   then show ?case by simp
       
   712 next
       
   713   case (insertI x A y)
       
   714   show ?case
   685   proof (cases "x = a")
   715   proof (cases "x = a")
   686     assume "x = a" with insertI show ?case by auto
   716     case True
       
   717     with insertI show ?thesis by auto
   687   next
   718   next
   688     assume "x \<noteq> a"
   719     case False
   689     then obtain y' where y: "y = f a y'" and y': "fold_graph f z (A - {a}) y'"
   720     then obtain y' where y: "y = f a y'" and y': "fold_graph f z (A - {a}) y'"
   690       using insertI by auto
   721       using insertI by auto
   691     have "f x y = f a (f x y')"
   722     have "f x y = f a (f x y')"
   692       unfolding y by (rule fun_left_comm)
   723       unfolding y by (rule fun_left_comm)
   693     moreover have "fold_graph f z (insert x A - {a}) (f x y')"
   724     moreover have "fold_graph f z (insert x A - {a}) (f x y')"
   694       using y' and \<open>x \<noteq> a\<close> and \<open>x \<notin> A\<close>
   725       using y' and \<open>x \<noteq> a\<close> and \<open>x \<notin> A\<close>
   695       by (simp add: insert_Diff_if fold_graph.insertI)
   726       by (simp add: insert_Diff_if fold_graph.insertI)
   696     ultimately show ?case by fast
   727     ultimately show ?thesis
   697   qed
   728       by fast
   698 qed simp
   729   qed
       
   730 qed
   699 
   731 
   700 lemma fold_graph_insertE:
   732 lemma fold_graph_insertE:
   701   assumes "fold_graph f z (insert x A) v" and "x \<notin> A"
   733   assumes "fold_graph f z (insert x A) v" and "x \<notin> A"
   702   obtains y where "v = f x y" and "fold_graph f z A y"
   734   obtains y where "v = f x y" and "fold_graph f z A y"
   703 using assms by (auto dest: fold_graph_insertE_aux [OF _ insertI1])
   735   using assms by (auto dest: fold_graph_insertE_aux [OF _ insertI1])
   704 
   736 
   705 lemma fold_graph_determ:
   737 lemma fold_graph_determ: "fold_graph f z A x \<Longrightarrow> fold_graph f z A y \<Longrightarrow> y = x"
   706   "fold_graph f z A x \<Longrightarrow> fold_graph f z A y \<Longrightarrow> y = x"
       
   707 proof (induct arbitrary: y set: fold_graph)
   738 proof (induct arbitrary: y set: fold_graph)
       
   739   case emptyI
       
   740   then show ?case by fast
       
   741 next
   708   case (insertI x A y v)
   742   case (insertI x A y v)
   709   from \<open>fold_graph f z (insert x A) v\<close> and \<open>x \<notin> A\<close>
   743   from \<open>fold_graph f z (insert x A) v\<close> and \<open>x \<notin> A\<close>
   710   obtain y' where "v = f x y'" and "fold_graph f z A y'"
   744   obtain y' where "v = f x y'" and "fold_graph f z A y'"
   711     by (rule fold_graph_insertE)
   745     by (rule fold_graph_insertE)
   712   from \<open>fold_graph f z A y'\<close> have "y' = y" by (rule insertI)
   746   from \<open>fold_graph f z A y'\<close> have "y' = y"
   713   with \<open>v = f x y'\<close> show "v = f x y" by simp
   747     by (rule insertI)
   714 qed fast
   748   with \<open>v = f x y'\<close> show "v = f x y"
   715 
   749     by simp
   716 lemma fold_equality:
   750 qed
   717   "fold_graph f z A y \<Longrightarrow> fold f z A = y"
   751 
       
   752 lemma fold_equality: "fold_graph f z A y \<Longrightarrow> fold f z A = y"
   718   by (cases "finite A") (auto simp add: fold_def intro: fold_graph_determ dest: fold_graph_finite)
   753   by (cases "finite A") (auto simp add: fold_def intro: fold_graph_determ dest: fold_graph_finite)
   719 
   754 
   720 lemma fold_graph_fold:
   755 lemma fold_graph_fold:
   721   assumes "finite A"
   756   assumes "finite A"
   722   shows "fold_graph f z A (fold f z A)"
   757   shows "fold_graph f z A (fold f z A)"
   723 proof -
   758 proof -
   724   from assms have "\<exists>x. fold_graph f z A x" by (rule finite_imp_fold_graph)
   759   from assms have "\<exists>x. fold_graph f z A x"
       
   760     by (rule finite_imp_fold_graph)
   725   moreover note fold_graph_determ
   761   moreover note fold_graph_determ
   726   ultimately have "\<exists>!x. fold_graph f z A x" by (rule ex_ex1I)
   762   ultimately have "\<exists>!x. fold_graph f z A x"
   727   then have "fold_graph f z A (The (fold_graph f z A))" by (rule theI')
   763     by (rule ex_ex1I)
   728   with assms show ?thesis by (simp add: fold_def)
   764   then have "fold_graph f z A (The (fold_graph f z A))"
       
   765     by (rule theI')
       
   766   with assms show ?thesis
       
   767     by (simp add: fold_def)
   729 qed
   768 qed
   730 
   769 
   731 text \<open>The base case for \<open>fold\<close>:\<close>
   770 text \<open>The base case for \<open>fold\<close>:\<close>
   732 
   771 
   733 lemma (in -) fold_infinite [simp]:
   772 lemma (in -) fold_infinite [simp]: "\<not> finite A \<Longrightarrow> fold f z A = z"
   734   assumes "\<not> finite A"
   773   by (auto simp: fold_def)
   735   shows "fold f z A = z"
   774 
   736   using assms by (auto simp add: fold_def)
   775 lemma (in -) fold_empty [simp]: "fold f z {} = z"
   737 
   776   by (auto simp: fold_def)
   738 lemma (in -) fold_empty [simp]:
   777 
   739   "fold f z {} = z"
   778 text \<open>The various recursion equations for @{const fold}:\<close>
   740   by (auto simp add: fold_def)
       
   741 
       
   742 text\<open>The various recursion equations for @{const fold}:\<close>
       
   743 
   779 
   744 lemma fold_insert [simp]:
   780 lemma fold_insert [simp]:
   745   assumes "finite A" and "x \<notin> A"
   781   assumes "finite A" and "x \<notin> A"
   746   shows "fold f z (insert x A) = f x (fold f z A)"
   782   shows "fold f z (insert x A) = f x (fold f z A)"
   747 proof (rule fold_equality)
   783 proof (rule fold_equality)
   748   fix z
   784   fix z
   749   from \<open>finite A\<close> have "fold_graph f z A (fold f z A)" by (rule fold_graph_fold)
   785   from \<open>finite A\<close> have "fold_graph f z A (fold f z A)"
   750   with \<open>x \<notin> A\<close> have "fold_graph f z (insert x A) (f x (fold f z A))" by (rule fold_graph.insertI)
   786     by (rule fold_graph_fold)
   751   then show "fold_graph f z (insert x A) (f x (fold f z A))" by simp
   787   with \<open>x \<notin> A\<close> have "fold_graph f z (insert x A) (f x (fold f z A))"
       
   788     by (rule fold_graph.insertI)
       
   789   then show "fold_graph f z (insert x A) (f x (fold f z A))"
       
   790     by simp
   752 qed
   791 qed
   753 
   792 
   754 declare (in -) empty_fold_graphE [rule del] fold_graph.intros [rule del]
   793 declare (in -) empty_fold_graphE [rule del] fold_graph.intros [rule del]
   755   \<comment> \<open>No more proofs involve these.\<close>
   794   \<comment> \<open>No more proofs involve these.\<close>
   756 
   795 
   757 lemma fold_fun_left_comm:
   796 lemma fold_fun_left_comm: "finite A \<Longrightarrow> f x (fold f z A) = fold f (f x z) A"
   758   "finite A \<Longrightarrow> f x (fold f z A) = fold f (f x z) A"
       
   759 proof (induct rule: finite_induct)
   797 proof (induct rule: finite_induct)
   760   case empty then show ?case by simp
   798   case empty
   761 next
   799   then show ?case by simp
   762   case (insert y A) then show ?case
   800 next
       
   801   case insert
       
   802   then show ?case
   763     by (simp add: fun_left_comm [of x])
   803     by (simp add: fun_left_comm [of x])
   764 qed
   804 qed
   765 
   805 
   766 lemma fold_insert2:
   806 lemma fold_insert2: "finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> fold f z (insert x A)  = fold f (f x z) A"
   767   "finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> fold f z (insert x A)  = fold f (f x z) A"
       
   768   by (simp add: fold_fun_left_comm)
   807   by (simp add: fold_fun_left_comm)
   769 
   808 
   770 lemma fold_rec:
   809 lemma fold_rec:
   771   assumes "finite A" and "x \<in> A"
   810   assumes "finite A" and "x \<in> A"
   772   shows "fold f z A = f x (fold f z (A - {x}))"
   811   shows "fold f z A = f x (fold f z (A - {x}))"
   773 proof -
   812 proof -
   774   have A: "A = insert x (A - {x})" using \<open>x \<in> A\<close> by blast
   813   have A: "A = insert x (A - {x})"
   775   then have "fold f z A = fold f z (insert x (A - {x}))" by simp
   814     using \<open>x \<in> A\<close> by blast
       
   815   then have "fold f z A = fold f z (insert x (A - {x}))"
       
   816     by simp
   776   also have "\<dots> = f x (fold f z (A - {x}))"
   817   also have "\<dots> = f x (fold f z (A - {x}))"
   777     by (rule fold_insert) (simp add: \<open>finite A\<close>)+
   818     by (rule fold_insert) (simp add: \<open>finite A\<close>)+
   778   finally show ?thesis .
   819   finally show ?thesis .
   779 qed
   820 qed
   780 
   821 
   781 lemma fold_insert_remove:
   822 lemma fold_insert_remove:
   782   assumes "finite A"
   823   assumes "finite A"
   783   shows "fold f z (insert x A) = f x (fold f z (A - {x}))"
   824   shows "fold f z (insert x A) = f x (fold f z (A - {x}))"
   784 proof -
   825 proof -
   785   from \<open>finite A\<close> have "finite (insert x A)" by auto
   826   from \<open>finite A\<close> have "finite (insert x A)"
   786   moreover have "x \<in> insert x A" by auto
   827     by auto
       
   828   moreover have "x \<in> insert x A"
       
   829     by auto
   787   ultimately have "fold f z (insert x A) = f x (fold f z (insert x A - {x}))"
   830   ultimately have "fold f z (insert x A) = f x (fold f z (insert x A - {x}))"
   788     by (rule fold_rec)
   831     by (rule fold_rec)
   789   then show ?thesis by simp
   832   then show ?thesis
       
   833     by simp
   790 qed
   834 qed
   791 
   835 
   792 lemma fold_set_union_disj:
   836 lemma fold_set_union_disj:
   793   assumes "finite A" "finite B" "A \<inter> B = {}"
   837   assumes "finite A" "finite B" "A \<inter> B = {}"
   794   shows "Finite_Set.fold f z (A \<union> B) = Finite_Set.fold f (Finite_Set.fold f z A) B"
   838   shows "Finite_Set.fold f z (A \<union> B) = Finite_Set.fold f (Finite_Set.fold f z A) B"
   795 using assms(2,1,3) by induction simp_all
   839   using assms(2,1,3) by induct simp_all
   796 
   840 
   797 end
   841 end
   798 
   842 
   799 text\<open>Other properties of @{const fold}:\<close>
   843 text \<open>Other properties of @{const fold}:\<close>
   800 
   844 
   801 lemma fold_image:
   845 lemma fold_image:
   802   assumes "inj_on g A"
   846   assumes "inj_on g A"
   803   shows "fold f z (g ` A) = fold (f \<circ> g) z A"
   847   shows "fold f z (g ` A) = fold (f \<circ> g) z A"
   804 proof (cases "finite A")
   848 proof (cases "finite A")
   805   case False with assms show ?thesis by (auto dest: finite_imageD simp add: fold_def)
   849   case False
       
   850   with assms show ?thesis
       
   851     by (auto dest: finite_imageD simp add: fold_def)
   806 next
   852 next
   807   case True
   853   case True
   808   have "fold_graph f z (g ` A) = fold_graph (f \<circ> g) z A"
   854   have "fold_graph f z (g ` A) = fold_graph (f \<circ> g) z A"
   809   proof
   855   proof
   810     fix w
   856     fix w
   811     show "fold_graph f z (g ` A) w \<longleftrightarrow> fold_graph (f \<circ> g) z A w" (is "?P \<longleftrightarrow> ?Q")
   857     show "fold_graph f z (g ` A) w \<longleftrightarrow> fold_graph (f \<circ> g) z A w" (is "?P \<longleftrightarrow> ?Q")
   812     proof
   858     proof
   813       assume ?P then show ?Q using assms
   859       assume ?P
       
   860       then show ?Q
       
   861         using assms
   814       proof (induct "g ` A" w arbitrary: A)
   862       proof (induct "g ` A" w arbitrary: A)
   815         case emptyI then show ?case by (auto intro: fold_graph.emptyI)
   863         case emptyI
       
   864         then show ?case by (auto intro: fold_graph.emptyI)
   816       next
   865       next
   817         case (insertI x A r B)
   866         case (insertI x A r B)
   818         from \<open>inj_on g B\<close> \<open>x \<notin> A\<close> \<open>insert x A = image g B\<close> obtain x' A' where
   867         from \<open>inj_on g B\<close> \<open>x \<notin> A\<close> \<open>insert x A = image g B\<close> obtain x' A'
   819           "x' \<notin> A'" and [simp]: "B = insert x' A'" "x = g x'" "A = g ` A'"
   868           where "x' \<notin> A'" and [simp]: "B = insert x' A'" "x = g x'" "A = g ` A'"
   820           by (rule inj_img_insertE)
   869           by (rule inj_img_insertE)
   821         from insertI.prems have "fold_graph (f o g) z A' r"
   870         from insertI.prems have "fold_graph (f \<circ> g) z A' r"
   822           by (auto intro: insertI.hyps)
   871           by (auto intro: insertI.hyps)
   823         with \<open>x' \<notin> A'\<close> have "fold_graph (f \<circ> g) z (insert x' A') ((f \<circ> g) x' r)"
   872         with \<open>x' \<notin> A'\<close> have "fold_graph (f \<circ> g) z (insert x' A') ((f \<circ> g) x' r)"
   824           by (rule fold_graph.insertI)
   873           by (rule fold_graph.insertI)
   825         then show ?case by simp
   874         then show ?case
       
   875           by simp
   826       qed
   876       qed
   827     next
   877     next
   828       assume ?Q then show ?P using assms
   878       assume ?Q
       
   879       then show ?P
       
   880         using assms
   829       proof induct
   881       proof induct
   830         case emptyI thus ?case by (auto intro: fold_graph.emptyI)
   882         case emptyI
       
   883         then show ?case
       
   884           by (auto intro: fold_graph.emptyI)
   831       next
   885       next
   832         case (insertI x A r)
   886         case (insertI x A r)
   833         from \<open>x \<notin> A\<close> insertI.prems have "g x \<notin> g ` A" by auto
   887         from \<open>x \<notin> A\<close> insertI.prems have "g x \<notin> g ` A"
   834         moreover from insertI have "fold_graph f z (g ` A) r" by simp
   888           by auto
       
   889         moreover from insertI have "fold_graph f z (g ` A) r"
       
   890           by simp
   835         ultimately have "fold_graph f z (insert (g x) (g ` A)) (f (g x) r)"
   891         ultimately have "fold_graph f z (insert (g x) (g ` A)) (f (g x) r)"
   836           by (rule fold_graph.insertI)
   892           by (rule fold_graph.insertI)
   837         then show ?case by simp
   893         then show ?case
       
   894           by simp
   838       qed
   895       qed
   839     qed
   896     qed
   840   qed
   897   qed
   841   with True assms show ?thesis by (auto simp add: fold_def)
   898   with True assms show ?thesis
       
   899     by (auto simp add: fold_def)
   842 qed
   900 qed
   843 
   901 
   844 lemma fold_cong:
   902 lemma fold_cong:
   845   assumes "comp_fun_commute f" "comp_fun_commute g"
   903   assumes "comp_fun_commute f" "comp_fun_commute g"
   846   assumes "finite A" and cong: "\<And>x. x \<in> A \<Longrightarrow> f x = g x"
   904     and "finite A"
       
   905     and cong: "\<And>x. x \<in> A \<Longrightarrow> f x = g x"
   847     and "s = t" and "A = B"
   906     and "s = t" and "A = B"
   848   shows "fold f s A = fold g t B"
   907   shows "fold f s A = fold g t B"
   849 proof -
   908 proof -
   850   have "fold f s A = fold g s A"  
   909   have "fold f s A = fold g s A"
   851   using \<open>finite A\<close> cong proof (induct A)
   910     using \<open>finite A\<close> cong
   852     case empty then show ?case by simp
   911   proof (induct A)
       
   912     case empty
       
   913     then show ?case by simp
   853   next
   914   next
   854     case (insert x A)
   915     case insert
   855     interpret f: comp_fun_commute f by (fact \<open>comp_fun_commute f\<close>)
   916     interpret f: comp_fun_commute f by (fact \<open>comp_fun_commute f\<close>)
   856     interpret g: comp_fun_commute g by (fact \<open>comp_fun_commute g\<close>)
   917     interpret g: comp_fun_commute g by (fact \<open>comp_fun_commute g\<close>)
   857     from insert show ?case by simp
   918     from insert show ?case by simp
   858   qed
   919   qed
   859   with assms show ?thesis by simp
   920   with assms show ?thesis by simp
   872 lemma fold_insert_idem:
   933 lemma fold_insert_idem:
   873   assumes fin: "finite A"
   934   assumes fin: "finite A"
   874   shows "fold f z (insert x A)  = f x (fold f z A)"
   935   shows "fold f z (insert x A)  = f x (fold f z A)"
   875 proof cases
   936 proof cases
   876   assume "x \<in> A"
   937   assume "x \<in> A"
   877   then obtain B where "A = insert x B" and "x \<notin> B" by (rule set_insert)
   938   then obtain B where "A = insert x B" and "x \<notin> B"
   878   then show ?thesis using assms by (simp add: comp_fun_idem fun_left_idem)
   939     by (rule set_insert)
   879 next
   940   then show ?thesis
   880   assume "x \<notin> A" then show ?thesis using assms by simp
   941     using assms by (simp add: comp_fun_idem fun_left_idem)
       
   942 next
       
   943   assume "x \<notin> A"
       
   944   then show ?thesis
       
   945     using assms by simp
   881 qed
   946 qed
   882 
   947 
   883 declare fold_insert [simp del] fold_insert_idem [simp]
   948 declare fold_insert [simp del] fold_insert_idem [simp]
   884 
   949 
   885 lemma fold_insert_idem2:
   950 lemma fold_insert_idem2: "finite A \<Longrightarrow> fold f z (insert x A) = fold f (f x z) A"
   886   "finite A \<Longrightarrow> fold f z (insert x A) = fold f (f x z) A"
       
   887   by (simp add: fold_fun_left_comm)
   951   by (simp add: fold_fun_left_comm)
   888 
   952 
   889 end
   953 end
   890 
   954 
   891 
   955 
   892 subsubsection \<open>Liftings to \<open>comp_fun_commute\<close> etc.\<close>
   956 subsubsection \<open>Liftings to \<open>comp_fun_commute\<close> etc.\<close>
   893 
   957 
   894 lemma (in comp_fun_commute) comp_comp_fun_commute:
   958 lemma (in comp_fun_commute) comp_comp_fun_commute: "comp_fun_commute (f \<circ> g)"
   895   "comp_fun_commute (f \<circ> g)"
   959   by standard (simp_all add: comp_fun_commute)
   896 proof
   960 
   897 qed (simp_all add: comp_fun_commute)
   961 lemma (in comp_fun_idem) comp_comp_fun_idem: "comp_fun_idem (f \<circ> g)"
   898 
       
   899 lemma (in comp_fun_idem) comp_comp_fun_idem:
       
   900   "comp_fun_idem (f \<circ> g)"
       
   901   by (rule comp_fun_idem.intro, rule comp_comp_fun_commute, unfold_locales)
   962   by (rule comp_fun_idem.intro, rule comp_comp_fun_commute, unfold_locales)
   902     (simp_all add: comp_fun_idem)
   963     (simp_all add: comp_fun_idem)
   903 
   964 
   904 lemma (in comp_fun_commute) comp_fun_commute_funpow:
   965 lemma (in comp_fun_commute) comp_fun_commute_funpow: "comp_fun_commute (\<lambda>x. f x ^^ g x)"
   905   "comp_fun_commute (\<lambda>x. f x ^^ g x)"
       
   906 proof
   966 proof
   907   fix y x
   967   show "f y ^^ g y \<circ> f x ^^ g x = f x ^^ g x \<circ> f y ^^ g y" for x y
   908   show "f y ^^ g y \<circ> f x ^^ g x = f x ^^ g x \<circ> f y ^^ g y"
       
   909   proof (cases "x = y")
   968   proof (cases "x = y")
   910     case True then show ?thesis by simp
   969     case True
       
   970     then show ?thesis by simp
   911   next
   971   next
   912     case False show ?thesis
   972     case False
       
   973     show ?thesis
   913     proof (induct "g x" arbitrary: g)
   974     proof (induct "g x" arbitrary: g)
   914       case 0 then show ?case by simp
   975       case 0
       
   976       then show ?case by simp
   915     next
   977     next
   916       case (Suc n g)
   978       case (Suc n g)
   917       have hyp1: "f y ^^ g y \<circ> f x = f x \<circ> f y ^^ g y"
   979       have hyp1: "f y ^^ g y \<circ> f x = f x \<circ> f y ^^ g y"
   918       proof (induct "g y" arbitrary: g)
   980       proof (induct "g y" arbitrary: g)
   919         case 0 then show ?case by simp
   981         case 0
       
   982         then show ?case by simp
   920       next
   983       next
   921         case (Suc n g)
   984         case (Suc n g)
   922         define h where "h z = g z - 1" for z
   985         define h where "h z = g z - 1" for z
   923         with Suc have "n = h y" by simp
   986         with Suc have "n = h y"
       
   987           by simp
   924         with Suc have hyp: "f y ^^ h y \<circ> f x = f x \<circ> f y ^^ h y"
   988         with Suc have hyp: "f y ^^ h y \<circ> f x = f x \<circ> f y ^^ h y"
   925           by auto
   989           by auto
   926         from Suc h_def have "g y = Suc (h y)" by simp
   990         from Suc h_def have "g y = Suc (h y)"
   927         then show ?case by (simp add: comp_assoc hyp)
   991           by simp
   928           (simp add: o_assoc comp_fun_commute)
   992         then show ?case
       
   993           by (simp add: comp_assoc hyp) (simp add: o_assoc comp_fun_commute)
   929       qed
   994       qed
   930       define h where "h z = (if z = x then g x - 1 else g z)" for z
   995       define h where "h z = (if z = x then g x - 1 else g z)" for z
   931       with Suc have "n = h x" by simp
   996       with Suc have "n = h x"
       
   997         by simp
   932       with Suc have "f y ^^ h y \<circ> f x ^^ h x = f x ^^ h x \<circ> f y ^^ h y"
   998       with Suc have "f y ^^ h y \<circ> f x ^^ h x = f x ^^ h x \<circ> f y ^^ h y"
   933         by auto
   999         by auto
   934       with False h_def have hyp2: "f y ^^ g y \<circ> f x ^^ h x = f x ^^ h x \<circ> f y ^^ g y" by simp
  1000       with False h_def have hyp2: "f y ^^ g y \<circ> f x ^^ h x = f x ^^ h x \<circ> f y ^^ g y"
   935       from Suc h_def have "g x = Suc (h x)" by simp
  1001         by simp
   936       then show ?case by (simp del: funpow.simps add: funpow_Suc_right o_assoc hyp2)
  1002       from Suc h_def have "g x = Suc (h x)"
   937         (simp add: comp_assoc hyp1)
  1003         by simp
       
  1004       then show ?case
       
  1005         by (simp del: funpow.simps add: funpow_Suc_right o_assoc hyp2) (simp add: comp_assoc hyp1)
   938     qed
  1006     qed
   939   qed
  1007   qed
   940 qed
  1008 qed
   941 
  1009 
   942 
  1010 
   943 subsubsection \<open>Expressing set operations via @{const fold}\<close>
  1011 subsubsection \<open>Expressing set operations via @{const fold}\<close>
   944 
  1012 
   945 lemma comp_fun_commute_const:
  1013 lemma comp_fun_commute_const: "comp_fun_commute (\<lambda>_. f)"
   946   "comp_fun_commute (\<lambda>_. f)"
  1014   by standard rule
   947 proof
  1015 
   948 qed rule
  1016 lemma comp_fun_idem_insert: "comp_fun_idem insert"
   949 
  1017   by standard auto
   950 lemma comp_fun_idem_insert:
  1018 
   951   "comp_fun_idem insert"
  1019 lemma comp_fun_idem_remove: "comp_fun_idem Set.remove"
   952 proof
  1020   by standard auto
   953 qed auto
  1021 
   954 
  1022 lemma (in semilattice_inf) comp_fun_idem_inf: "comp_fun_idem inf"
   955 lemma comp_fun_idem_remove:
  1023   by standard (auto simp add: inf_left_commute)
   956   "comp_fun_idem Set.remove"
  1024 
   957 proof
  1025 lemma (in semilattice_sup) comp_fun_idem_sup: "comp_fun_idem sup"
   958 qed auto
  1026   by standard (auto simp add: sup_left_commute)
   959 
       
   960 lemma (in semilattice_inf) comp_fun_idem_inf:
       
   961   "comp_fun_idem inf"
       
   962 proof
       
   963 qed (auto simp add: inf_left_commute)
       
   964 
       
   965 lemma (in semilattice_sup) comp_fun_idem_sup:
       
   966   "comp_fun_idem sup"
       
   967 proof
       
   968 qed (auto simp add: sup_left_commute)
       
   969 
  1027 
   970 lemma union_fold_insert:
  1028 lemma union_fold_insert:
   971   assumes "finite A"
  1029   assumes "finite A"
   972   shows "A \<union> B = fold insert B A"
  1030   shows "A \<union> B = fold insert B A"
   973 proof -
  1031 proof -
   974   interpret comp_fun_idem insert by (fact comp_fun_idem_insert)
  1032   interpret comp_fun_idem insert
   975   from \<open>finite A\<close> show ?thesis by (induct A arbitrary: B) simp_all
  1033     by (fact comp_fun_idem_insert)
       
  1034   from \<open>finite A\<close> show ?thesis
       
  1035     by (induct A arbitrary: B) simp_all
   976 qed
  1036 qed
   977 
  1037 
   978 lemma minus_fold_remove:
  1038 lemma minus_fold_remove:
   979   assumes "finite A"
  1039   assumes "finite A"
   980   shows "B - A = fold Set.remove B A"
  1040   shows "B - A = fold Set.remove B A"
   981 proof -
  1041 proof -
   982   interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove)
  1042   interpret comp_fun_idem Set.remove
   983   from \<open>finite A\<close> have "fold Set.remove B A = B - A" by (induct A arbitrary: B) auto
  1043     by (fact comp_fun_idem_remove)
       
  1044   from \<open>finite A\<close> have "fold Set.remove B A = B - A"
       
  1045     by (induct A arbitrary: B) auto
   984   then show ?thesis ..
  1046   then show ?thesis ..
   985 qed
  1047 qed
   986 
  1048 
   987 lemma comp_fun_commute_filter_fold:
  1049 lemma comp_fun_commute_filter_fold:
   988   "comp_fun_commute (\<lambda>x A'. if P x then Set.insert x A' else A')"
  1050   "comp_fun_commute (\<lambda>x A'. if P x then Set.insert x A' else A')"
   989 proof - 
  1051 proof -
   990   interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert)
  1052   interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert)
   991   show ?thesis by standard (auto simp: fun_eq_iff)
  1053   show ?thesis by standard (auto simp: fun_eq_iff)
   992 qed
  1054 qed
   993 
  1055 
   994 lemma Set_filter_fold:
  1056 lemma Set_filter_fold:
   995   assumes "finite A"
  1057   assumes "finite A"
   996   shows "Set.filter P A = fold (\<lambda>x A'. if P x then Set.insert x A' else A') {} A"
  1058   shows "Set.filter P A = fold (\<lambda>x A'. if P x then Set.insert x A' else A') {} A"
   997 using assms
  1059   using assms
   998 by (induct A) 
  1060   by induct
   999   (auto simp add: Set.filter_def comp_fun_commute.fold_insert[OF comp_fun_commute_filter_fold])
  1061     (auto simp add: Set.filter_def comp_fun_commute.fold_insert[OF comp_fun_commute_filter_fold])
  1000 
  1062 
  1001 lemma inter_Set_filter:     
  1063 lemma inter_Set_filter:
  1002   assumes "finite B"
  1064   assumes "finite B"
  1003   shows "A \<inter> B = Set.filter (\<lambda>x. x \<in> A) B"
  1065   shows "A \<inter> B = Set.filter (\<lambda>x. x \<in> A) B"
  1004 using assms 
  1066   using assms
  1005 by (induct B) (auto simp: Set.filter_def)
  1067   by induct (auto simp: Set.filter_def)
  1006 
  1068 
  1007 lemma image_fold_insert:
  1069 lemma image_fold_insert:
  1008   assumes "finite A"
  1070   assumes "finite A"
  1009   shows "image f A = fold (\<lambda>k A. Set.insert (f k) A) {} A"
  1071   shows "image f A = fold (\<lambda>k A. Set.insert (f k) A) {} A"
  1010 using assms
  1072 proof -
  1011 proof -
  1073   interpret comp_fun_commute "\<lambda>k A. Set.insert (f k) A"
  1012   interpret comp_fun_commute "\<lambda>k A. Set.insert (f k) A" by standard auto
  1074     by standard auto
  1013   show ?thesis using assms by (induct A) auto
  1075   show ?thesis
       
  1076     using assms by (induct A) auto
  1014 qed
  1077 qed
  1015 
  1078 
  1016 lemma Ball_fold:
  1079 lemma Ball_fold:
  1017   assumes "finite A"
  1080   assumes "finite A"
  1018   shows "Ball A P = fold (\<lambda>k s. s \<and> P k) True A"
  1081   shows "Ball A P = fold (\<lambda>k s. s \<and> P k) True A"
  1019 using assms
  1082 proof -
  1020 proof -
  1083   interpret comp_fun_commute "\<lambda>k s. s \<and> P k"
  1021   interpret comp_fun_commute "\<lambda>k s. s \<and> P k" by standard auto
  1084     by standard auto
  1022   show ?thesis using assms by (induct A) auto
  1085   show ?thesis
       
  1086     using assms by (induct A) auto
  1023 qed
  1087 qed
  1024 
  1088 
  1025 lemma Bex_fold:
  1089 lemma Bex_fold:
  1026   assumes "finite A"
  1090   assumes "finite A"
  1027   shows "Bex A P = fold (\<lambda>k s. s \<or> P k) False A"
  1091   shows "Bex A P = fold (\<lambda>k s. s \<or> P k) False A"
  1028 using assms
  1092 proof -
  1029 proof -
  1093   interpret comp_fun_commute "\<lambda>k s. s \<or> P k"
  1030   interpret comp_fun_commute "\<lambda>k s. s \<or> P k" by standard auto
  1094     by standard auto
  1031   show ?thesis using assms by (induct A) auto
  1095   show ?thesis
  1032 qed
  1096     using assms by (induct A) auto
  1033 
  1097 qed
  1034 lemma comp_fun_commute_Pow_fold: 
  1098 
  1035   "comp_fun_commute (\<lambda>x A. A \<union> Set.insert x ` A)" 
  1099 lemma comp_fun_commute_Pow_fold: "comp_fun_commute (\<lambda>x A. A \<union> Set.insert x ` A)"
  1036   by (clarsimp simp: fun_eq_iff comp_fun_commute_def) blast
  1100   by (clarsimp simp: fun_eq_iff comp_fun_commute_def) blast
  1037 
  1101 
  1038 lemma Pow_fold:
  1102 lemma Pow_fold:
  1039   assumes "finite A"
  1103   assumes "finite A"
  1040   shows "Pow A = fold (\<lambda>x A. A \<union> Set.insert x ` A) {{}} A"
  1104   shows "Pow A = fold (\<lambda>x A. A \<union> Set.insert x ` A) {{}} A"
  1041 using assms
  1105 proof -
  1042 proof -
  1106   interpret comp_fun_commute "\<lambda>x A. A \<union> Set.insert x ` A"
  1043   interpret comp_fun_commute "\<lambda>x A. A \<union> Set.insert x ` A" by (rule comp_fun_commute_Pow_fold)
  1107     by (rule comp_fun_commute_Pow_fold)
  1044   show ?thesis using assms by (induct A) (auto simp: Pow_insert)
  1108   show ?thesis
       
  1109     using assms by (induct A) (auto simp: Pow_insert)
  1045 qed
  1110 qed
  1046 
  1111 
  1047 lemma fold_union_pair:
  1112 lemma fold_union_pair:
  1048   assumes "finite B"
  1113   assumes "finite B"
  1049   shows "(\<Union>y\<in>B. {(x, y)}) \<union> A = fold (\<lambda>y. Set.insert (x, y)) A B"
  1114   shows "(\<Union>y\<in>B. {(x, y)}) \<union> A = fold (\<lambda>y. Set.insert (x, y)) A B"
  1050 proof -
  1115 proof -
  1051   interpret comp_fun_commute "\<lambda>y. Set.insert (x, y)" by standard auto
  1116   interpret comp_fun_commute "\<lambda>y. Set.insert (x, y)"
  1052   show ?thesis using assms  by (induct B arbitrary: A) simp_all
  1117     by standard auto
  1053 qed
  1118   show ?thesis
  1054 
  1119     using assms by (induct arbitrary: A) simp_all
  1055 lemma comp_fun_commute_product_fold: 
  1120 qed
  1056   assumes "finite B"
  1121 
  1057   shows "comp_fun_commute (\<lambda>x z. fold (\<lambda>y. Set.insert (x, y)) z B)" 
  1122 lemma comp_fun_commute_product_fold:
  1058   by standard (auto simp: fold_union_pair[symmetric] assms)
  1123   "finite B \<Longrightarrow> comp_fun_commute (\<lambda>x z. fold (\<lambda>y. Set.insert (x, y)) z B)"
       
  1124   by standard (auto simp: fold_union_pair [symmetric])
  1059 
  1125 
  1060 lemma product_fold:
  1126 lemma product_fold:
  1061   assumes "finite A"
  1127   assumes "finite A" "finite B"
  1062   assumes "finite B"
       
  1063   shows "A \<times> B = fold (\<lambda>x z. fold (\<lambda>y. Set.insert (x, y)) z B) {} A"
  1128   shows "A \<times> B = fold (\<lambda>x z. fold (\<lambda>y. Set.insert (x, y)) z B) {} A"
  1064 using assms unfolding Sigma_def 
  1129   using assms unfolding Sigma_def
  1065 by (induct A) 
  1130   by (induct A)
  1066   (simp_all add: comp_fun_commute.fold_insert[OF comp_fun_commute_product_fold] fold_union_pair)
  1131     (simp_all add: comp_fun_commute.fold_insert[OF comp_fun_commute_product_fold] fold_union_pair)
  1067 
       
  1068 
  1132 
  1069 context complete_lattice
  1133 context complete_lattice
  1070 begin
  1134 begin
  1071 
  1135 
  1072 lemma inf_Inf_fold_inf:
  1136 lemma inf_Inf_fold_inf:
  1073   assumes "finite A"
  1137   assumes "finite A"
  1074   shows "inf (Inf A) B = fold inf B A"
  1138   shows "inf (Inf A) B = fold inf B A"
  1075 proof -
  1139 proof -
  1076   interpret comp_fun_idem inf by (fact comp_fun_idem_inf)
  1140   interpret comp_fun_idem inf
  1077   from \<open>finite A\<close> fold_fun_left_comm show ?thesis by (induct A arbitrary: B)
  1141     by (fact comp_fun_idem_inf)
  1078     (simp_all add: inf_commute fun_eq_iff)
  1142   from \<open>finite A\<close> fold_fun_left_comm show ?thesis
       
  1143     by (induct A arbitrary: B) (simp_all add: inf_commute fun_eq_iff)
  1079 qed
  1144 qed
  1080 
  1145 
  1081 lemma sup_Sup_fold_sup:
  1146 lemma sup_Sup_fold_sup:
  1082   assumes "finite A"
  1147   assumes "finite A"
  1083   shows "sup (Sup A) B = fold sup B A"
  1148   shows "sup (Sup A) B = fold sup B A"
  1084 proof -
  1149 proof -
  1085   interpret comp_fun_idem sup by (fact comp_fun_idem_sup)
  1150   interpret comp_fun_idem sup
  1086   from \<open>finite A\<close> fold_fun_left_comm show ?thesis by (induct A arbitrary: B)
  1151     by (fact comp_fun_idem_sup)
  1087     (simp_all add: sup_commute fun_eq_iff)
  1152   from \<open>finite A\<close> fold_fun_left_comm show ?thesis
  1088 qed
  1153     by (induct A arbitrary: B) (simp_all add: sup_commute fun_eq_iff)
  1089 
  1154 qed
  1090 lemma Inf_fold_inf:
  1155 
  1091   assumes "finite A"
  1156 lemma Inf_fold_inf: "finite A \<Longrightarrow> Inf A = fold inf top A"
  1092   shows "Inf A = fold inf top A"
  1157   using inf_Inf_fold_inf [of A top] by (simp add: inf_absorb2)
  1093   using assms inf_Inf_fold_inf [of A top] by (simp add: inf_absorb2)
  1158 
  1094 
  1159 lemma Sup_fold_sup: "finite A \<Longrightarrow> Sup A = fold sup bot A"
  1095 lemma Sup_fold_sup:
  1160   using sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2)
  1096   assumes "finite A"
       
  1097   shows "Sup A = fold sup bot A"
       
  1098   using assms sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2)
       
  1099 
  1161 
  1100 lemma inf_INF_fold_inf:
  1162 lemma inf_INF_fold_inf:
  1101   assumes "finite A"
  1163   assumes "finite A"
  1102   shows "inf B (INFIMUM A f) = fold (inf \<circ> f) B A" (is "?inf = ?fold") 
  1164   shows "inf B (INFIMUM A f) = fold (inf \<circ> f) B A" (is "?inf = ?fold")
  1103 proof (rule sym)
  1165 proof -
  1104   interpret comp_fun_idem inf by (fact comp_fun_idem_inf)
  1166   interpret comp_fun_idem inf by (fact comp_fun_idem_inf)
  1105   interpret comp_fun_idem "inf \<circ> f" by (fact comp_comp_fun_idem)
  1167   interpret comp_fun_idem "inf \<circ> f" by (fact comp_comp_fun_idem)
  1106   from \<open>finite A\<close> show "?fold = ?inf"
  1168   from \<open>finite A\<close> have "?fold = ?inf"
  1107     by (induct A arbitrary: B)
  1169     by (induct A arbitrary: B) (simp_all add: inf_left_commute)
  1108       (simp_all add: inf_left_commute)
  1170   then show ?thesis ..
  1109 qed
  1171 qed
  1110 
  1172 
  1111 lemma sup_SUP_fold_sup:
  1173 lemma sup_SUP_fold_sup:
  1112   assumes "finite A"
  1174   assumes "finite A"
  1113   shows "sup B (SUPREMUM A f) = fold (sup \<circ> f) B A" (is "?sup = ?fold") 
  1175   shows "sup B (SUPREMUM A f) = fold (sup \<circ> f) B A" (is "?sup = ?fold")
  1114 proof (rule sym)
  1176 proof -
  1115   interpret comp_fun_idem sup by (fact comp_fun_idem_sup)
  1177   interpret comp_fun_idem sup by (fact comp_fun_idem_sup)
  1116   interpret comp_fun_idem "sup \<circ> f" by (fact comp_comp_fun_idem)
  1178   interpret comp_fun_idem "sup \<circ> f" by (fact comp_comp_fun_idem)
  1117   from \<open>finite A\<close> show "?fold = ?sup"
  1179   from \<open>finite A\<close> have "?fold = ?sup"
  1118     by (induct A arbitrary: B)
  1180     by (induct A arbitrary: B) (simp_all add: sup_left_commute)
  1119       (simp_all add: sup_left_commute)
  1181   then show ?thesis ..
  1120 qed
  1182 qed
  1121 
  1183 
  1122 lemma INF_fold_inf:
  1184 lemma INF_fold_inf: "finite A \<Longrightarrow> INFIMUM A f = fold (inf \<circ> f) top A"
  1123   assumes "finite A"
  1185   using inf_INF_fold_inf [of A top] by simp
  1124   shows "INFIMUM A f = fold (inf \<circ> f) top A"
  1186 
  1125   using assms inf_INF_fold_inf [of A top] by simp
  1187 lemma SUP_fold_sup: "finite A \<Longrightarrow> SUPREMUM A f = fold (sup \<circ> f) bot A"
  1126 
  1188   using sup_SUP_fold_sup [of A bot] by simp
  1127 lemma SUP_fold_sup:
       
  1128   assumes "finite A"
       
  1129   shows "SUPREMUM A f = fold (sup \<circ> f) bot A"
       
  1130   using assms sup_SUP_fold_sup [of A bot] by simp
       
  1131 
  1189 
  1132 end
  1190 end
  1133 
  1191 
  1134 
  1192 
  1135 subsection \<open>Locales as mini-packages for fold operations\<close>
  1193 subsection \<open>Locales as mini-packages for fold operations\<close>
  1144 
  1202 
  1145 interpretation fold?: comp_fun_commute f
  1203 interpretation fold?: comp_fun_commute f
  1146   by standard (insert comp_fun_commute, simp add: fun_eq_iff)
  1204   by standard (insert comp_fun_commute, simp add: fun_eq_iff)
  1147 
  1205 
  1148 definition F :: "'a set \<Rightarrow> 'b"
  1206 definition F :: "'a set \<Rightarrow> 'b"
  1149 where
  1207   where eq_fold: "F A = fold f z A"
  1150   eq_fold: "F A = fold f z A"
       
  1151 
  1208 
  1152 lemma empty [simp]:"F {} = z"
  1209 lemma empty [simp]:"F {} = z"
  1153   by (simp add: eq_fold)
  1210   by (simp add: eq_fold)
  1154 
  1211 
  1155 lemma infinite [simp]: "\<not> finite A \<Longrightarrow> F A = z"
  1212 lemma infinite [simp]: "\<not> finite A \<Longrightarrow> F A = z"
  1156   by (simp add: eq_fold)
  1213   by (simp add: eq_fold)
  1157  
  1214 
  1158 lemma insert [simp]:
  1215 lemma insert [simp]:
  1159   assumes "finite A" and "x \<notin> A"
  1216   assumes "finite A" and "x \<notin> A"
  1160   shows "F (insert x A) = f x (F A)"
  1217   shows "F (insert x A) = f x (F A)"
  1161 proof -
  1218 proof -
  1162   from fold_insert assms
  1219   from fold_insert assms
  1163   have "fold f z (insert x A) = f x (fold f z A)" by simp
  1220   have "fold f z (insert x A) = f x (fold f z A)" by simp
  1164   with \<open>finite A\<close> show ?thesis by (simp add: eq_fold fun_eq_iff)
  1221   with \<open>finite A\<close> show ?thesis by (simp add: eq_fold fun_eq_iff)
  1165 qed
  1222 qed
  1166  
  1223 
  1167 lemma remove:
  1224 lemma remove:
  1168   assumes "finite A" and "x \<in> A"
  1225   assumes "finite A" and "x \<in> A"
  1169   shows "F A = f x (F (A - {x}))"
  1226   shows "F A = f x (F (A - {x}))"
  1170 proof -
  1227 proof -
  1171   from \<open>x \<in> A\<close> obtain B where A: "A = insert x B" and "x \<notin> B"
  1228   from \<open>x \<in> A\<close> obtain B where A: "A = insert x B" and "x \<notin> B"
  1172     by (auto dest: mk_disjoint_insert)
  1229     by (auto dest: mk_disjoint_insert)
  1173   moreover from \<open>finite A\<close> A have "finite B" by simp
  1230   moreover from \<open>finite A\<close> A have "finite B" by simp
  1174   ultimately show ?thesis by simp
  1231   ultimately show ?thesis by simp
  1175 qed
  1232 qed
  1176 
  1233 
  1177 lemma insert_remove:
  1234 lemma insert_remove: "finite A \<Longrightarrow> F (insert x A) = f x (F (A - {x}))"
  1178   assumes "finite A"
  1235   by (cases "x \<in> A") (simp_all add: remove insert_absorb)
  1179   shows "F (insert x A) = f x (F (A - {x}))"
       
  1180   using assms by (cases "x \<in> A") (simp_all add: remove insert_absorb)
       
  1181 
  1236 
  1182 end
  1237 end
  1183 
  1238 
  1184 
  1239 
  1185 subsubsection \<open>With idempotency\<close>
  1240 subsubsection \<open>With idempotency\<close>
  1207 
  1262 
  1208 subsection \<open>Finite cardinality\<close>
  1263 subsection \<open>Finite cardinality\<close>
  1209 
  1264 
  1210 text \<open>
  1265 text \<open>
  1211   The traditional definition
  1266   The traditional definition
  1212   @{prop "card A \<equiv> LEAST n. EX f. A = {f i | i. i < n}"}
  1267   @{prop "card A \<equiv> LEAST n. \<exists>f. A = {f i |i. i < n}"}
  1213   is ugly to work with.
  1268   is ugly to work with.
  1214   But now that we have @{const fold} things are easy:
  1269   But now that we have @{const fold} things are easy:
  1215 \<close>
  1270 \<close>
  1216 
  1271 
  1217 global_interpretation card: folding "\<lambda>_. Suc" 0
  1272 global_interpretation card: folding "\<lambda>_. Suc" 0
  1218   defines card = "folding.F (\<lambda>_. Suc) 0"
  1273   defines card = "folding.F (\<lambda>_. Suc) 0"
  1219   by standard rule
  1274   by standard rule
  1220 
  1275 
  1221 lemma card_infinite:
  1276 lemma card_infinite: "\<not> finite A \<Longrightarrow> card A = 0"
  1222   "\<not> finite A \<Longrightarrow> card A = 0"
       
  1223   by (fact card.infinite)
  1277   by (fact card.infinite)
  1224 
  1278 
  1225 lemma card_empty:
  1279 lemma card_empty: "card {} = 0"
  1226   "card {} = 0"
       
  1227   by (fact card.empty)
  1280   by (fact card.empty)
  1228 
  1281 
  1229 lemma card_insert_disjoint:
  1282 lemma card_insert_disjoint: "finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> card (insert x A) = Suc (card A)"
  1230   "finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> card (insert x A) = Suc (card A)"
       
  1231   by (fact card.insert)
  1283   by (fact card.insert)
  1232 
  1284 
  1233 lemma card_insert_if:
  1285 lemma card_insert_if: "finite A \<Longrightarrow> card (insert x A) = (if x \<in> A then card A else Suc (card A))"
  1234   "finite A \<Longrightarrow> card (insert x A) = (if x \<in> A then card A else Suc (card A))"
       
  1235   by auto (simp add: card.insert_remove card.remove)
  1286   by auto (simp add: card.insert_remove card.remove)
  1236 
  1287 
  1237 lemma card_ge_0_finite:
  1288 lemma card_ge_0_finite: "card A > 0 \<Longrightarrow> finite A"
  1238   "card A > 0 \<Longrightarrow> finite A"
       
  1239   by (rule ccontr) simp
  1289   by (rule ccontr) simp
  1240 
  1290 
  1241 lemma card_0_eq [simp]:
  1291 lemma card_0_eq [simp]: "finite A \<Longrightarrow> card A = 0 \<longleftrightarrow> A = {}"
  1242   "finite A \<Longrightarrow> card A = 0 \<longleftrightarrow> A = {}"
       
  1243   by (auto dest: mk_disjoint_insert)
  1292   by (auto dest: mk_disjoint_insert)
  1244 
  1293 
  1245 lemma finite_UNIV_card_ge_0:
  1294 lemma finite_UNIV_card_ge_0: "finite (UNIV :: 'a set) \<Longrightarrow> card (UNIV :: 'a set) > 0"
  1246   "finite (UNIV :: 'a set) \<Longrightarrow> card (UNIV :: 'a set) > 0"
       
  1247   by (rule ccontr) simp
  1295   by (rule ccontr) simp
  1248 
  1296 
  1249 lemma card_eq_0_iff:
  1297 lemma card_eq_0_iff: "card A = 0 \<longleftrightarrow> A = {} \<or> \<not> finite A"
  1250   "card A = 0 \<longleftrightarrow> A = {} \<or> \<not> finite A"
       
  1251   by auto
  1298   by auto
  1252 
  1299 
  1253 lemma card_range_greater_zero:
  1300 lemma card_range_greater_zero: "finite (range f) \<Longrightarrow> card (range f) > 0"
  1254   "finite (range f) \<Longrightarrow> card (range f) > 0"
       
  1255   by (rule ccontr) (simp add: card_eq_0_iff)
  1301   by (rule ccontr) (simp add: card_eq_0_iff)
  1256 
  1302 
  1257 lemma card_gt_0_iff:
  1303 lemma card_gt_0_iff: "0 < card A \<longleftrightarrow> A \<noteq> {} \<and> finite A"
  1258   "0 < card A \<longleftrightarrow> A \<noteq> {} \<and> finite A"
  1304   by (simp add: neq0_conv [symmetric] card_eq_0_iff)
  1259   by (simp add: neq0_conv [symmetric] card_eq_0_iff) 
  1305 
  1260 
  1306 lemma card_Suc_Diff1: "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> Suc (card (A - {x})) = card A"
  1261 lemma card_Suc_Diff1:
  1307   apply (rule insert_Diff [THEN subst, where t = A])
  1262   "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> Suc (card (A - {x})) = card A"
  1308   apply assumption
  1263 apply(rule_tac t = A in insert_Diff [THEN subst], assumption)
  1309   apply (simp del: insert_Diff_single)
  1264 apply(simp del:insert_Diff_single)
  1310   done
  1265 done
  1311 
  1266 
  1312 lemma card_insert_le_m1: "n > 0 \<Longrightarrow> card y \<le> n - 1 \<Longrightarrow> card (insert x y) \<le> n"
  1267 lemma card_insert_le_m1: "n>0 \<Longrightarrow> card y \<le> n-1 \<Longrightarrow> card (insert x y) \<le> n"
       
  1268   apply (cases "finite y")
  1313   apply (cases "finite y")
  1269   apply (cases "x \<in> y")
  1314   apply (cases "x \<in> y")
  1270   apply (auto simp: insert_absorb)
  1315   apply (auto simp: insert_absorb)
  1271   done
  1316   done
  1272 
  1317 
  1273 lemma card_Diff_singleton:
  1318 lemma card_Diff_singleton: "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> card (A - {x}) = card A - 1"
  1274   "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> card (A - {x}) = card A - 1"
       
  1275   by (simp add: card_Suc_Diff1 [symmetric])
  1319   by (simp add: card_Suc_Diff1 [symmetric])
  1276 
  1320 
  1277 lemma card_Diff_singleton_if:
  1321 lemma card_Diff_singleton_if:
  1278   "finite A \<Longrightarrow> card (A - {x}) = (if x \<in> A then card A - 1 else card A)"
  1322   "finite A \<Longrightarrow> card (A - {x}) = (if x \<in> A then card A - 1 else card A)"
  1279   by (simp add: card_Diff_singleton)
  1323   by (simp add: card_Diff_singleton)
  1280 
  1324 
  1281 lemma card_Diff_insert[simp]:
  1325 lemma card_Diff_insert[simp]:
  1282   assumes "finite A" and "a \<in> A" and "a \<notin> B"
  1326   assumes "finite A" and "a \<in> A" and "a \<notin> B"
  1283   shows "card (A - insert a B) = card (A - B) - 1"
  1327   shows "card (A - insert a B) = card (A - B) - 1"
  1284 proof -
  1328 proof -
  1285   have "A - insert a B = (A - B) - {a}" using assms by blast
  1329   have "A - insert a B = (A - B) - {a}"
  1286   then show ?thesis using assms by(simp add: card_Diff_singleton)
  1330     using assms by blast
  1287 qed
  1331   then show ?thesis
  1288 
  1332     using assms by (simp add: card_Diff_singleton)
  1289 lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))"
  1333 qed
       
  1334 
       
  1335 lemma card_insert: "finite A \<Longrightarrow> card (insert x A) = Suc (card (A - {x}))"
  1290   by (fact card.insert_remove)
  1336   by (fact card.insert_remove)
  1291 
  1337 
  1292 lemma card_insert_le: "finite A ==> card A <= card (insert x A)"
  1338 lemma card_insert_le: "finite A \<Longrightarrow> card A \<le> card (insert x A)"
  1293 by (simp add: card_insert_if)
  1339   by (simp add: card_insert_if)
  1294 
  1340 
  1295 lemma card_Collect_less_nat[simp]: "card{i::nat. i < n} = n"
  1341 lemma card_Collect_less_nat[simp]: "card {i::nat. i < n} = n"
  1296 by (induct n) (simp_all add:less_Suc_eq Collect_disj_eq)
  1342   by (induct n) (simp_all add:less_Suc_eq Collect_disj_eq)
  1297 
  1343 
  1298 lemma card_Collect_le_nat[simp]: "card{i::nat. i <= n} = Suc n"
  1344 lemma card_Collect_le_nat[simp]: "card {i::nat. i \<le> n} = Suc n"
  1299 using card_Collect_less_nat[of "Suc n"] by(simp add: less_Suc_eq_le)
  1345   using card_Collect_less_nat[of "Suc n"] by (simp add: less_Suc_eq_le)
  1300 
  1346 
  1301 lemma card_mono:
  1347 lemma card_mono:
  1302   assumes "finite B" and "A \<subseteq> B"
  1348   assumes "finite B" and "A \<subseteq> B"
  1303   shows "card A \<le> card B"
  1349   shows "card A \<le> card B"
  1304 proof -
  1350 proof -
  1305   from assms have "finite A" by (auto intro: finite_subset)
  1351   from assms have "finite A"
  1306   then show ?thesis using assms proof (induct A arbitrary: B)
  1352     by (auto intro: finite_subset)
  1307     case empty then show ?case by simp
  1353   then show ?thesis
       
  1354     using assms
       
  1355   proof (induct A arbitrary: B)
       
  1356     case empty
       
  1357     then show ?case by simp
  1308   next
  1358   next
  1309     case (insert x A)
  1359     case (insert x A)
  1310     then have "x \<in> B" by simp
  1360     then have "x \<in> B"
  1311     from insert have "A \<subseteq> B - {x}" and "finite (B - {x})" by auto
  1361       by simp
  1312     with insert.hyps have "card A \<le> card (B - {x})" by auto
  1362     from insert have "A \<subseteq> B - {x}" and "finite (B - {x})"
  1313     with \<open>finite A\<close> \<open>x \<notin> A\<close> \<open>finite B\<close> \<open>x \<in> B\<close> show ?case by simp (simp only: card.remove)
  1363       by auto
  1314   qed
  1364     with insert.hyps have "card A \<le> card (B - {x})"
  1315 qed
  1365       by auto
  1316 
  1366     with \<open>finite A\<close> \<open>x \<notin> A\<close> \<open>finite B\<close> \<open>x \<in> B\<close> show ?case
  1317 lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)"
  1367       by simp (simp only: card.remove)
  1318 apply (induct rule: finite_induct)
  1368   qed
  1319 apply simp
  1369 qed
  1320 apply clarify
  1370 
  1321 apply (subgoal_tac "finite A & A - {x} <= F")
  1371 lemma card_seteq: "finite B \<Longrightarrow> (\<And>A. A \<subseteq> B \<Longrightarrow> card B \<le> card A \<Longrightarrow> A = B)"
  1322  prefer 2 apply (blast intro: finite_subset, atomize)
  1372   apply (induct rule: finite_induct)
  1323 apply (drule_tac x = "A - {x}" in spec)
  1373   apply simp
  1324 apply (simp add: card_Diff_singleton_if split add: if_split_asm)
  1374   apply clarify
  1325 apply (case_tac "card A", auto)
  1375   apply (subgoal_tac "finite A \<and> A - {x} \<subseteq> F")
  1326 done
  1376    prefer 2 apply (blast intro: finite_subset, atomize)
  1327 
  1377   apply (drule_tac x = "A - {x}" in spec)
  1328 lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B"
  1378   apply (simp add: card_Diff_singleton_if split add: if_split_asm)
  1329 apply (simp add: psubset_eq linorder_not_le [symmetric])
  1379   apply (case_tac "card A", auto)
  1330 apply (blast dest: card_seteq)
  1380   done
  1331 done
  1381 
       
  1382 lemma psubset_card_mono: "finite B \<Longrightarrow> A < B \<Longrightarrow> card A < card B"
       
  1383   apply (simp add: psubset_eq linorder_not_le [symmetric])
       
  1384   apply (blast dest: card_seteq)
       
  1385   done
  1332 
  1386 
  1333 lemma card_Un_Int:
  1387 lemma card_Un_Int:
  1334   assumes "finite A" and "finite B"
  1388   assumes "finite A" "finite B"
  1335   shows "card A + card B = card (A \<union> B) + card (A \<inter> B)"
  1389   shows "card A + card B = card (A \<union> B) + card (A \<inter> B)"
  1336 using assms proof (induct A)
  1390   using assms
  1337   case empty then show ?case by simp
  1391 proof (induct A)
  1338 next
  1392   case empty
  1339  case (insert x A) then show ?case
  1393   then show ?case by simp
       
  1394 next
       
  1395   case insert
       
  1396   then show ?case
  1340     by (auto simp add: insert_absorb Int_insert_left)
  1397     by (auto simp add: insert_absorb Int_insert_left)
  1341 qed
  1398 qed
  1342 
  1399 
  1343 lemma card_Un_disjoint:
  1400 lemma card_Un_disjoint: "finite A \<Longrightarrow> finite B \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> card (A \<union> B) = card A + card B"
  1344   assumes "finite A" and "finite B"
  1401   using card_Un_Int [of A B] by simp
  1345   assumes "A \<inter> B = {}"
       
  1346   shows "card (A \<union> B) = card A + card B"
       
  1347 using assms card_Un_Int [of A B] by simp
       
  1348 
  1402 
  1349 lemma card_Un_le: "card (A \<union> B) \<le> card A + card B"
  1403 lemma card_Un_le: "card (A \<union> B) \<le> card A + card B"
  1350 apply(cases "finite A")
  1404   apply (cases "finite A")
  1351  apply(cases "finite B")
  1405    apply (cases "finite B")
  1352   using le_iff_add card_Un_Int apply blast
  1406     using le_iff_add card_Un_Int apply blast
  1353  apply simp
  1407    apply simp
  1354 apply simp
  1408   apply simp
  1355 done
  1409   done
  1356 
  1410 
  1357 lemma card_Diff_subset:
  1411 lemma card_Diff_subset:
  1358   assumes "finite B" and "B \<subseteq> A"
  1412   assumes "finite B"
       
  1413     and "B \<subseteq> A"
  1359   shows "card (A - B) = card A - card B"
  1414   shows "card (A - B) = card A - card B"
  1360 proof (cases "finite A")
  1415 proof (cases "finite A")
  1361   case False with assms show ?thesis by simp
  1416   case False
  1362 next
  1417   with assms show ?thesis
  1363   case True with assms show ?thesis by (induct B arbitrary: A) simp_all
  1418     by simp
       
  1419 next
       
  1420   case True
       
  1421   with assms show ?thesis
       
  1422     by (induct B arbitrary: A) simp_all
  1364 qed
  1423 qed
  1365 
  1424 
  1366 lemma card_Diff_subset_Int:
  1425 lemma card_Diff_subset_Int:
  1367   assumes AB: "finite (A \<inter> B)" shows "card (A - B) = card A - card (A \<inter> B)"
  1426   assumes "finite (A \<inter> B)"
       
  1427   shows "card (A - B) = card A - card (A \<inter> B)"
  1368 proof -
  1428 proof -
  1369   have "A - B = A - A \<inter> B" by auto
  1429   have "A - B = A - A \<inter> B" by auto
  1370   thus ?thesis
  1430   with assms show ?thesis
  1371     by (simp add: card_Diff_subset AB) 
  1431     by (simp add: card_Diff_subset)
  1372 qed
  1432 qed
  1373 
  1433 
  1374 lemma diff_card_le_card_Diff:
  1434 lemma diff_card_le_card_Diff:
  1375 assumes "finite B" shows "card A - card B \<le> card(A - B)"
  1435   assumes "finite B"
  1376 proof-
  1436   shows "card A - card B \<le> card (A - B)"
       
  1437 proof -
  1377   have "card A - card B \<le> card A - card (A \<inter> B)"
  1438   have "card A - card B \<le> card A - card (A \<inter> B)"
  1378     using card_mono[OF assms Int_lower2, of A] by arith
  1439     using card_mono[OF assms Int_lower2, of A] by arith
  1379   also have "\<dots> = card(A-B)" using assms by(simp add: card_Diff_subset_Int)
  1440   also have "\<dots> = card (A - B)"
       
  1441     using assms by (simp add: card_Diff_subset_Int)
  1380   finally show ?thesis .
  1442   finally show ?thesis .
  1381 qed
  1443 qed
  1382 
  1444 
  1383 lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A"
  1445 lemma card_Diff1_less: "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> card (A - {x}) < card A"
  1384 apply (rule Suc_less_SucD)
  1446   by (rule Suc_less_SucD) (simp add: card_Suc_Diff1 del: card_Diff_insert)
  1385 apply (simp add: card_Suc_Diff1 del:card_Diff_insert)
  1447 
  1386 done
  1448 lemma card_Diff2_less: "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> card (A - {x} - {y}) < card A"
  1387 
  1449   apply (cases "x = y")
  1388 lemma card_Diff2_less:
  1450    apply (simp add: card_Diff1_less del:card_Diff_insert)
  1389   "finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A"
  1451   apply (rule less_trans)
  1390 apply (case_tac "x = y")
  1452    prefer 2 apply (auto intro!: card_Diff1_less simp del: card_Diff_insert)
  1391  apply (simp add: card_Diff1_less del:card_Diff_insert)
  1453   done
  1392 apply (rule less_trans)
  1454 
  1393  prefer 2 apply (auto intro!: card_Diff1_less simp del:card_Diff_insert)
  1455 lemma card_Diff1_le: "finite A \<Longrightarrow> card (A - {x}) \<le> card A"
  1394 done
  1456   by (cases "x \<in> A") (simp_all add: card_Diff1_less less_imp_le)
  1395 
  1457 
  1396 lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A"
  1458 lemma card_psubset: "finite B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> card A < card B \<Longrightarrow> A < B"
  1397 apply (case_tac "x : A")
  1459   by (erule psubsetI) blast
  1398  apply (simp_all add: card_Diff1_less less_imp_le)
       
  1399 done
       
  1400 
       
  1401 lemma card_psubset: "finite B ==> A \<subseteq> B ==> card A < card B ==> A < B"
       
  1402 by (erule psubsetI, blast)
       
  1403 
  1460 
  1404 lemma card_le_inj:
  1461 lemma card_le_inj:
  1405   assumes fA: "finite A"
  1462   assumes fA: "finite A"
  1406     and fB: "finite B"
  1463     and fB: "finite B"
  1407     and c: "card A \<le> card B"
  1464     and c: "card A \<le> card B"
  1411   case empty
  1468   case empty
  1412   then show ?case by simp
  1469   then show ?case by simp
  1413 next
  1470 next
  1414   case (insert x s t)
  1471   case (insert x s t)
  1415   then show ?case
  1472   then show ?case
  1416   proof (induct rule: finite_induct[OF "insert.prems"(1)])
  1473   proof (induct rule: finite_induct [OF insert.prems(1)])
  1417     case 1
  1474     case 1
  1418     then show ?case by simp
  1475     then show ?case by simp
  1419   next
  1476   next
  1420     case (2 y t)
  1477     case (2 y t)
  1421     from "2.prems"(1,2,5) "2.hyps"(1,2) have cst: "card s \<le> card t"
  1478     from "2.prems"(1,2,5) "2.hyps"(1,2) have cst: "card s \<le> card t"
  1452   with AB show "A = B"
  1509   with AB show "A = B"
  1453     by blast
  1510     by blast
  1454 qed
  1511 qed
  1455 
  1512 
  1456 lemma insert_partition:
  1513 lemma insert_partition:
  1457   "\<lbrakk> x \<notin> F; \<forall>c1 \<in> insert x F. \<forall>c2 \<in> insert x F. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {} \<rbrakk>
  1514   "x \<notin> F \<Longrightarrow> \<forall>c1 \<in> insert x F. \<forall>c2 \<in> insert x F. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {} \<Longrightarrow> x \<inter> \<Union>F = {}"
  1458   \<Longrightarrow> x \<inter> \<Union>F = {}"
  1515   by auto
  1459 by auto
  1516 
  1460 
  1517 lemma finite_psubset_induct [consumes 1, case_names psubset]:
  1461 lemma finite_psubset_induct[consumes 1, case_names psubset]:
  1518   assumes finite: "finite A"
  1462   assumes fin: "finite A" 
  1519     and major: "\<And>A. finite A \<Longrightarrow> (\<And>B. B \<subset> A \<Longrightarrow> P B) \<Longrightarrow> P A"
  1463   and     major: "\<And>A. finite A \<Longrightarrow> (\<And>B. B \<subset> A \<Longrightarrow> P B) \<Longrightarrow> P A" 
       
  1464   shows "P A"
  1520   shows "P A"
  1465 using fin
  1521   using finite
  1466 proof (induct A taking: card rule: measure_induct_rule)
  1522 proof (induct A taking: card rule: measure_induct_rule)
  1467   case (less A)
  1523   case (less A)
  1468   have fin: "finite A" by fact
  1524   have fin: "finite A" by fact
  1469   have ih: "\<And>B. \<lbrakk>card B < card A; finite B\<rbrakk> \<Longrightarrow> P B" by fact
  1525   have ih: "card B < card A \<Longrightarrow> finite B \<Longrightarrow> P B" for B by fact
  1470   { fix B 
  1526   have "P B" if "B \<subset> A" for B
  1471     assume asm: "B \<subset> A"
  1527   proof -
  1472     from asm have "card B < card A" using psubset_card_mono fin by blast
  1528     from that have "card B < card A"
       
  1529       using psubset_card_mono fin by blast
  1473     moreover
  1530     moreover
  1474     from asm have "B \<subseteq> A" by auto
  1531     from that have "B \<subseteq> A"
  1475     then have "finite B" using fin finite_subset by blast
  1532       by auto
  1476     ultimately 
  1533     then have "finite B"
  1477     have "P B" using ih by simp
  1534       using fin finite_subset by blast
  1478   }
  1535     ultimately show ?thesis using ih by simp
       
  1536   qed
  1479   with fin show "P A" using major by blast
  1537   with fin show "P A" using major by blast
  1480 qed
  1538 qed
  1481 
  1539 
  1482 lemma finite_induct_select[consumes 1, case_names empty select]:
  1540 lemma finite_induct_select [consumes 1, case_names empty select]:
  1483   assumes "finite S"
  1541   assumes "finite S"
  1484   assumes "P {}"
  1542     and "P {}"
  1485   assumes select: "\<And>T. T \<subset> S \<Longrightarrow> P T \<Longrightarrow> \<exists>s\<in>S - T. P (insert s T)"
  1543     and select: "\<And>T. T \<subset> S \<Longrightarrow> P T \<Longrightarrow> \<exists>s\<in>S - T. P (insert s T)"
  1486   shows "P S"
  1544   shows "P S"
  1487 proof -
  1545 proof -
  1488   have "0 \<le> card S" by simp
  1546   have "0 \<le> card S" by simp
  1489   then have "\<exists>T \<subseteq> S. card T = card S \<and> P T"
  1547   then have "\<exists>T \<subseteq> S. card T = card S \<and> P T"
  1490   proof (induct rule: dec_induct)
  1548   proof (induct rule: dec_induct)
  1491     case base with \<open>P {}\<close> show ?case
  1549     case base with \<open>P {}\<close>
       
  1550     show ?case
  1492       by (intro exI[of _ "{}"]) auto
  1551       by (intro exI[of _ "{}"]) auto
  1493   next
  1552   next
  1494     case (step n)
  1553     case (step n)
  1495     then obtain T where T: "T \<subseteq> S" "card T = n" "P T"
  1554     then obtain T where T: "T \<subseteq> S" "card T = n" "P T"
  1496       by auto
  1555       by auto
  1504   with \<open>finite S\<close> show "P S"
  1563   with \<open>finite S\<close> show "P S"
  1505     by (auto dest: card_subset_eq)
  1564     by (auto dest: card_subset_eq)
  1506 qed
  1565 qed
  1507 
  1566 
  1508 lemma remove_induct [case_names empty infinite remove]:
  1567 lemma remove_induct [case_names empty infinite remove]:
  1509   assumes empty: "P ({} :: 'a set)" and infinite: "\<not>finite B \<Longrightarrow> P B"
  1568   assumes empty: "P ({} :: 'a set)"
  1510       and remove: "\<And>A. finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> P (A - {x})) \<Longrightarrow> P A"
  1569     and infinite: "\<not> finite B \<Longrightarrow> P B"
       
  1570     and remove: "\<And>A. finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> P (A - {x})) \<Longrightarrow> P A"
  1511   shows "P B"
  1571   shows "P B"
  1512 proof (cases "finite B")
  1572 proof (cases "finite B")
  1513   assume "\<not>finite B"
  1573   assume "\<not>finite B"
  1514   thus ?thesis by (rule infinite)
  1574   then show ?thesis by (rule infinite)
  1515 next
  1575 next
  1516   define A where "A = B"
  1576   define A where "A = B"
  1517   assume "finite B"
  1577   assume "finite B"
  1518   hence "finite A" "A \<subseteq> B" by (simp_all add: A_def)
  1578   then have "finite A" "A \<subseteq> B"
  1519   thus "P A"
  1579     by (simp_all add: A_def)
  1520   proof (induction "card A" arbitrary: A)
  1580   then show "P A"
       
  1581   proof (induct "card A" arbitrary: A)
  1521     case 0
  1582     case 0
  1522     hence "A = {}" by auto
  1583     then have "A = {}" by auto
  1523     with empty show ?case by simp
  1584     with empty show ?case by simp
  1524   next
  1585   next
  1525     case (Suc n A)
  1586     case (Suc n A)
  1526     from \<open>A \<subseteq> B\<close> and \<open>finite B\<close> have "finite A" by (rule finite_subset)
  1587     from \<open>A \<subseteq> B\<close> and \<open>finite B\<close> have "finite A"
       
  1588       by (rule finite_subset)
  1527     moreover from Suc.hyps have "A \<noteq> {}" by auto
  1589     moreover from Suc.hyps have "A \<noteq> {}" by auto
  1528     moreover note \<open>A \<subseteq> B\<close>
  1590     moreover note \<open>A \<subseteq> B\<close>
  1529     moreover have "P (A - {x})" if x: "x \<in> A" for x
  1591     moreover have "P (A - {x})" if x: "x \<in> A" for x
  1530       using x Suc.prems \<open>Suc n = card A\<close> by (intro Suc) auto
  1592       using x Suc.prems \<open>Suc n = card A\<close> by (intro Suc) auto
  1531     ultimately show ?case by (rule remove)
  1593     ultimately show ?case by (rule remove)
  1532   qed
  1594   qed
  1533 qed
  1595 qed
  1534 
  1596 
  1535 lemma finite_remove_induct [consumes 1, case_names empty remove]:
  1597 lemma finite_remove_induct [consumes 1, case_names empty remove]:
  1536   assumes finite: "finite B" and empty: "P ({} :: 'a set)" 
  1598   fixes P :: "'a set \<Rightarrow> bool"
  1537       and rm: "\<And>A. finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> P (A - {x})) \<Longrightarrow> P A"
  1599   assumes finite: "finite B"
       
  1600     and empty: "P {}"
       
  1601     and rm: "\<And>A. finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> P (A - {x})) \<Longrightarrow> P A"
  1538   defines "B' \<equiv> B"
  1602   defines "B' \<equiv> B"
  1539   shows   "P B'"
  1603   shows "P B'"
  1540   by (induction B' rule: remove_induct) (simp_all add: assms)
  1604   by (induct B' rule: remove_induct) (simp_all add: assms)
  1541 
  1605 
  1542 
  1606 
  1543 text\<open>main cardinality theorem\<close>
  1607 text \<open>Main cardinality theorem.\<close>
  1544 lemma card_partition [rule_format]:
  1608 lemma card_partition [rule_format]:
  1545   "finite C ==>
  1609   "finite C \<Longrightarrow> finite (\<Union>C) \<Longrightarrow> (\<forall>c\<in>C. card c = k) \<Longrightarrow>
  1546      finite (\<Union>C) -->
  1610     (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {}) \<Longrightarrow>
  1547      (\<forall>c\<in>C. card c = k) -->
  1611     k * card C = card (\<Union>C)"
  1548      (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 --> c1 \<inter> c2 = {}) -->
  1612   apply (induct rule: finite_induct)
  1549      k * card(C) = card (\<Union>C)"
  1613   apply simp
  1550 apply (erule finite_induct, simp)
  1614   apply (simp add: card_Un_disjoint insert_partition finite_subset [of _ "\<Union>(insert x F)"])
  1551 apply (simp add: card_Un_disjoint insert_partition 
  1615   done
  1552        finite_subset [of _ "\<Union>(insert x F)"])
       
  1553 done
       
  1554 
  1616 
  1555 lemma card_eq_UNIV_imp_eq_UNIV:
  1617 lemma card_eq_UNIV_imp_eq_UNIV:
  1556   assumes fin: "finite (UNIV :: 'a set)"
  1618   assumes fin: "finite (UNIV :: 'a set)"
  1557   and card: "card A = card (UNIV :: 'a set)"
  1619     and card: "card A = card (UNIV :: 'a set)"
  1558   shows "A = (UNIV :: 'a set)"
  1620   shows "A = (UNIV :: 'a set)"
  1559 proof
  1621 proof
  1560   show "A \<subseteq> UNIV" by simp
  1622   show "A \<subseteq> UNIV" by simp
  1561   show "UNIV \<subseteq> A"
  1623   show "UNIV \<subseteq> A"
  1562   proof
  1624   proof
  1563     fix x
  1625     show "x \<in> A" for x
  1564     show "x \<in> A"
       
  1565     proof (rule ccontr)
  1626     proof (rule ccontr)
  1566       assume "x \<notin> A"
  1627       assume "x \<notin> A"
  1567       then have "A \<subset> UNIV" by auto
  1628       then have "A \<subset> UNIV" by auto
  1568       with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono)
  1629       with fin have "card A < card (UNIV :: 'a set)"
       
  1630         by (fact psubset_card_mono)
  1569       with card show False by simp
  1631       with card show False by simp
  1570     qed
  1632     qed
  1571   qed
  1633   qed
  1572 qed
  1634 qed
  1573 
  1635 
  1574 text\<open>The form of a finite set of given cardinality\<close>
  1636 text \<open>The form of a finite set of given cardinality\<close>
  1575 
  1637 
  1576 lemma card_eq_SucD:
  1638 lemma card_eq_SucD:
  1577 assumes "card A = Suc k"
  1639   assumes "card A = Suc k"
  1578 shows "\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={})"
  1640   shows "\<exists>b B. A = insert b B \<and> b \<notin> B \<and> card B = k \<and> (k = 0 \<longrightarrow> B = {})"
  1579 proof -
  1641 proof -
  1580   have fin: "finite A" using assms by (auto intro: ccontr)
  1642   have fin: "finite A"
  1581   moreover have "card A \<noteq> 0" using assms by auto
  1643     using assms by (auto intro: ccontr)
  1582   ultimately obtain b where b: "b \<in> A" by auto
  1644   moreover have "card A \<noteq> 0"
       
  1645     using assms by auto
       
  1646   ultimately obtain b where b: "b \<in> A"
       
  1647     by auto
  1583   show ?thesis
  1648   show ?thesis
  1584   proof (intro exI conjI)
  1649   proof (intro exI conjI)
  1585     show "A = insert b (A-{b})" using b by blast
  1650     show "A = insert b (A - {b})"
  1586     show "b \<notin> A - {b}" by blast
  1651       using b by blast
       
  1652     show "b \<notin> A - {b}"
       
  1653       by blast
  1587     show "card (A - {b}) = k" and "k = 0 \<longrightarrow> A - {b} = {}"
  1654     show "card (A - {b}) = k" and "k = 0 \<longrightarrow> A - {b} = {}"
  1588       using assms b fin by(fastforce dest:mk_disjoint_insert)+
  1655       using assms b fin by(fastforce dest:mk_disjoint_insert)+
  1589   qed
  1656   qed
  1590 qed
  1657 qed
  1591 
  1658 
  1592 lemma card_Suc_eq:
  1659 lemma card_Suc_eq:
  1593   "(card A = Suc k) =
  1660   "card A = Suc k \<longleftrightarrow>
  1594    (\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={}))"
  1661     (\<exists>b B. A = insert b B \<and> b \<notin> B \<and> card B = k \<and> (k = 0 \<longrightarrow> B = {}))"
  1595  apply(auto elim!: card_eq_SucD)
  1662   apply (auto elim!: card_eq_SucD)
  1596  apply(subst card.insert)
  1663   apply (subst card.insert)
  1597  apply(auto simp add: intro:ccontr)
  1664   apply (auto simp add: intro:ccontr)
  1598  done
  1665   done
  1599 
  1666 
  1600 lemma card_1_singletonE:
  1667 lemma card_1_singletonE:
  1601     assumes "card A = 1" obtains x where "A = {x}"
  1668   assumes "card A = 1"
       
  1669   obtains x where "A = {x}"
  1602   using assms by (auto simp: card_Suc_eq)
  1670   using assms by (auto simp: card_Suc_eq)
  1603 
  1671 
  1604 lemma is_singleton_altdef: "is_singleton A \<longleftrightarrow> card A = 1"
  1672 lemma is_singleton_altdef: "is_singleton A \<longleftrightarrow> card A = 1"
  1605   unfolding is_singleton_def
  1673   unfolding is_singleton_def
  1606   by (auto elim!: card_1_singletonE is_singletonE simp del: One_nat_def)
  1674   by (auto elim!: card_1_singletonE is_singletonE simp del: One_nat_def)
  1607 
  1675 
  1608 lemma card_le_Suc_iff: "finite A \<Longrightarrow>
  1676 lemma card_le_Suc_iff:
  1609   Suc n \<le> card A = (\<exists>a B. A = insert a B \<and> a \<notin> B \<and> n \<le> card B \<and> finite B)"
  1677   "finite A \<Longrightarrow> Suc n \<le> card A = (\<exists>a B. A = insert a B \<and> a \<notin> B \<and> n \<le> card B \<and> finite B)"
  1610 by (fastforce simp: card_Suc_eq less_eq_nat.simps(2) insert_eq_iff
  1678   by (fastforce simp: card_Suc_eq less_eq_nat.simps(2) insert_eq_iff
  1611   dest: subset_singletonD split: nat.splits if_splits)
  1679     dest: subset_singletonD split: nat.splits if_splits)
  1612 
  1680 
  1613 lemma finite_fun_UNIVD2:
  1681 lemma finite_fun_UNIVD2:
  1614   assumes fin: "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
  1682   assumes fin: "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
  1615   shows "finite (UNIV :: 'b set)"
  1683   shows "finite (UNIV :: 'b set)"
  1616 proof -
  1684 proof -
  1617   from fin have "\<And>arbitrary. finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary))"
  1685   from fin have "finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary))" for arbitrary
  1618     by (rule finite_imageI)
  1686     by (rule finite_imageI)
  1619   moreover have "\<And>arbitrary. UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary)"
  1687   moreover have "UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary)" for arbitrary
  1620     by (rule UNIV_eq_I) auto
  1688     by (rule UNIV_eq_I) auto
  1621   ultimately show "finite (UNIV :: 'b set)" by simp
  1689   ultimately show "finite (UNIV :: 'b set)"
       
  1690     by simp
  1622 qed
  1691 qed
  1623 
  1692 
  1624 lemma card_UNIV_unit [simp]: "card (UNIV :: unit set) = 1"
  1693 lemma card_UNIV_unit [simp]: "card (UNIV :: unit set) = 1"
  1625   unfolding UNIV_unit by simp
  1694   unfolding UNIV_unit by simp
  1626 
  1695 
  1627 lemma infinite_arbitrarily_large:
  1696 lemma infinite_arbitrarily_large:
  1628   assumes "\<not> finite A"
  1697   assumes "\<not> finite A"
  1629   shows "\<exists>B. finite B \<and> card B = n \<and> B \<subseteq> A"
  1698   shows "\<exists>B. finite B \<and> card B = n \<and> B \<subseteq> A"
  1630 proof (induction n)
  1699 proof (induction n)
  1631   case 0 show ?case by (intro exI[of _ "{}"]) auto
  1700   case 0
  1632 next 
  1701   show ?case by (intro exI[of _ "{}"]) auto
       
  1702 next
  1633   case (Suc n)
  1703   case (Suc n)
  1634   then guess B .. note B = this
  1704   then obtain B where B: "finite B \<and> card B = n \<and> B \<subseteq> A" ..
  1635   with \<open>\<not> finite A\<close> have "A \<noteq> B" by auto
  1705   with \<open>\<not> finite A\<close> have "A \<noteq> B" by auto
  1636   with B have "B \<subset> A" by auto
  1706   with B have "B \<subset> A" by auto
  1637   hence "\<exists>x. x \<in> A - B" by (elim psubset_imp_ex_mem)
  1707   then have "\<exists>x. x \<in> A - B"
  1638   then guess x .. note x = this
  1708     by (elim psubset_imp_ex_mem)
       
  1709   then obtain x where x: "x \<in> A - B" ..
  1639   with B have "finite (insert x B) \<and> card (insert x B) = Suc n \<and> insert x B \<subseteq> A"
  1710   with B have "finite (insert x B) \<and> card (insert x B) = Suc n \<and> insert x B \<subseteq> A"
  1640     by auto
  1711     by auto
  1641   thus "\<exists>B. finite B \<and> card B = Suc n \<and> B \<subseteq> A" ..
  1712   then show "\<exists>B. finite B \<and> card B = Suc n \<and> B \<subseteq> A" ..
  1642 qed
  1713 qed
       
  1714 
  1643 
  1715 
  1644 subsubsection \<open>Cardinality of image\<close>
  1716 subsubsection \<open>Cardinality of image\<close>
  1645 
  1717 
  1646 lemma card_image_le: "finite A ==> card (f ` A) \<le> card A"
  1718 lemma card_image_le: "finite A \<Longrightarrow> card (f ` A) \<le> card A"
  1647   by (induct rule: finite_induct) (simp_all add: le_SucI card_insert_if)
  1719   by (induct rule: finite_induct) (simp_all add: le_SucI card_insert_if)
  1648 
  1720 
  1649 lemma card_image:
  1721 lemma card_image:
  1650   assumes "inj_on f A"
  1722   assumes "inj_on f A"
  1651   shows "card (f ` A) = card A"
  1723   shows "card (f ` A) = card A"
  1652 proof (cases "finite A")
  1724 proof (cases "finite A")
  1653   case True then show ?thesis using assms by (induct A) simp_all
  1725   case True
  1654 next
  1726   then show ?thesis
  1655   case False then have "\<not> finite (f ` A)" using assms by (auto dest: finite_imageD)
  1727     using assms by (induct A) simp_all
  1656   with False show ?thesis by simp
  1728 next
       
  1729   case False
       
  1730   then have "\<not> finite (f ` A)"
       
  1731     using assms by (auto dest: finite_imageD)
       
  1732   with False show ?thesis
       
  1733     by simp
  1657 qed
  1734 qed
  1658 
  1735 
  1659 lemma bij_betw_same_card: "bij_betw f A B \<Longrightarrow> card A = card B"
  1736 lemma bij_betw_same_card: "bij_betw f A B \<Longrightarrow> card A = card B"
  1660 by(auto simp: card_image bij_betw_def)
  1737   by(auto simp: card_image bij_betw_def)
  1661 
  1738 
  1662 lemma endo_inj_surj: "finite A ==> f ` A \<subseteq> A ==> inj_on f A ==> f ` A = A"
  1739 lemma endo_inj_surj: "finite A \<Longrightarrow> f ` A \<subseteq> A \<Longrightarrow> inj_on f A \<Longrightarrow> f ` A = A"
  1663 by (simp add: card_seteq card_image)
  1740   by (simp add: card_seteq card_image)
  1664 
  1741 
  1665 lemma eq_card_imp_inj_on:
  1742 lemma eq_card_imp_inj_on:
  1666   assumes "finite A" "card(f ` A) = card A" shows "inj_on f A"
  1743   assumes "finite A" "card(f ` A) = card A"
  1667 using assms
  1744   shows "inj_on f A"
       
  1745   using assms
  1668 proof (induct rule:finite_induct)
  1746 proof (induct rule:finite_induct)
  1669   case empty show ?case by simp
  1747   case empty
       
  1748   show ?case by simp
  1670 next
  1749 next
  1671   case (insert x A)
  1750   case (insert x A)
  1672   then show ?case using card_image_le [of A f]
  1751   then show ?case
  1673     by (simp add: card_insert_if split: if_splits)
  1752     using card_image_le [of A f] by (simp add: card_insert_if split: if_splits)
  1674 qed
  1753 qed
  1675 
  1754 
  1676 lemma inj_on_iff_eq_card: "finite A \<Longrightarrow> inj_on f A \<longleftrightarrow> card(f ` A) = card A"
  1755 lemma inj_on_iff_eq_card: "finite A \<Longrightarrow> inj_on f A \<longleftrightarrow> card (f ` A) = card A"
  1677   by (blast intro: card_image eq_card_imp_inj_on)
  1756   by (blast intro: card_image eq_card_imp_inj_on)
  1678 
  1757 
  1679 lemma card_inj_on_le:
  1758 lemma card_inj_on_le:
  1680   assumes "inj_on f A" "f ` A \<subseteq> B" "finite B" shows "card A \<le> card B"
  1759   assumes "inj_on f A" "f ` A \<subseteq> B" "finite B"
  1681 proof -
  1760   shows "card A \<le> card B"
  1682   have "finite A" using assms
  1761 proof -
  1683     by (blast intro: finite_imageD dest: finite_subset)
  1762   have "finite A"
  1684   then show ?thesis using assms 
  1763     using assms by (blast intro: finite_imageD dest: finite_subset)
  1685    by (force intro: card_mono simp: card_image [symmetric])
  1764   then show ?thesis
       
  1765     using assms by (force intro: card_mono simp: card_image [symmetric])
  1686 qed
  1766 qed
  1687 
  1767 
  1688 lemma surj_card_le: "finite A \<Longrightarrow> B \<subseteq> f ` A \<Longrightarrow> card B \<le> card A"
  1768 lemma surj_card_le: "finite A \<Longrightarrow> B \<subseteq> f ` A \<Longrightarrow> card B \<le> card A"
  1689   by (blast intro: card_image_le card_mono le_trans)
  1769   by (blast intro: card_image_le card_mono le_trans)
  1690 
  1770 
  1691 lemma card_bij_eq:
  1771 lemma card_bij_eq:
  1692   "[|inj_on f A; f ` A \<subseteq> B; inj_on g B; g ` B \<subseteq> A;
  1772   "inj_on f A \<Longrightarrow> f ` A \<subseteq> B \<Longrightarrow> inj_on g B \<Longrightarrow> g ` B \<subseteq> A \<Longrightarrow> finite A \<Longrightarrow> finite B
  1693      finite A; finite B |] ==> card A = card B"
  1773     \<Longrightarrow> card A = card B"
  1694 by (auto intro: le_antisym card_inj_on_le)
  1774   by (auto intro: le_antisym card_inj_on_le)
  1695 
  1775 
  1696 lemma bij_betw_finite:
  1776 lemma bij_betw_finite: "bij_betw f A B \<Longrightarrow> finite A \<longleftrightarrow> finite B"
  1697   assumes "bij_betw f A B"
  1777   unfolding bij_betw_def using finite_imageD [of f A] by auto
  1698   shows "finite A \<longleftrightarrow> finite B"
  1778 
  1699 using assms unfolding bij_betw_def
  1779 lemma inj_on_finite: "inj_on f A \<Longrightarrow> f ` A \<le> B \<Longrightarrow> finite B \<Longrightarrow> finite A"
  1700 using finite_imageD[of f A] by auto
  1780   using finite_imageD finite_subset by blast
  1701 
  1781 
  1702 lemma inj_on_finite:
  1782 lemma card_vimage_inj: "inj f \<Longrightarrow> A \<subseteq> range f \<Longrightarrow> card (f -` A) = card A"
  1703 assumes "inj_on f A" "f ` A \<le> B" "finite B"
  1783   by (auto 4 3 simp: subset_image_iff inj_vimage_image_eq
  1704 shows "finite A"
  1784       intro: card_image[symmetric, OF subset_inj_on])
  1705 using assms finite_imageD finite_subset by blast
  1785 
  1706 
       
  1707 lemma card_vimage_inj: "\<lbrakk> inj f; A \<subseteq> range f \<rbrakk> \<Longrightarrow> card (f -` A) = card A"
       
  1708 by(auto 4 3 simp add: subset_image_iff inj_vimage_image_eq intro: card_image[symmetric, OF subset_inj_on])
       
  1709 
  1786 
  1710 subsubsection \<open>Pigeonhole Principles\<close>
  1787 subsubsection \<open>Pigeonhole Principles\<close>
  1711 
  1788 
  1712 lemma pigeonhole: "card A > card(f ` A) \<Longrightarrow> ~ inj_on f A "
  1789 lemma pigeonhole: "card A > card (f ` A) \<Longrightarrow> \<not> inj_on f A "
  1713 by (auto dest: card_image less_irrefl_nat)
  1790   by (auto dest: card_image less_irrefl_nat)
  1714 
  1791 
  1715 lemma pigeonhole_infinite:
  1792 lemma pigeonhole_infinite:
  1716 assumes  "~ finite A" and "finite(f`A)"
  1793   assumes "\<not> finite A" and "finite (f`A)"
  1717 shows "EX a0:A. ~finite{a:A. f a = f a0}"
  1794   shows "\<exists>a0\<in>A. \<not> finite {a\<in>A. f a = f a0}"
  1718 proof -
  1795   using assms(2,1)
  1719   have "finite(f`A) \<Longrightarrow> ~ finite A \<Longrightarrow> EX a0:A. ~finite{a:A. f a = f a0}"
  1796 proof (induct "f`A" arbitrary: A rule: finite_induct)
  1720   proof(induct "f`A" arbitrary: A rule: finite_induct)
  1797   case empty
  1721     case empty thus ?case by simp
  1798   then show ?case by simp
       
  1799 next
       
  1800   case (insert b F)
       
  1801   show ?case
       
  1802   proof (cases "finite {a\<in>A. f a = b}")
       
  1803     case True
       
  1804     with \<open>\<not> finite A\<close> have "\<not> finite (A - {a\<in>A. f a = b})"
       
  1805       by simp
       
  1806     also have "A - {a\<in>A. f a = b} = {a\<in>A. f a \<noteq> b}"
       
  1807       by blast
       
  1808     finally have "\<not> finite {a\<in>A. f a \<noteq> b}" .
       
  1809     from insert(3)[OF _ this] insert(2,4) show ?thesis
       
  1810       by simp (blast intro: rev_finite_subset)
  1722   next
  1811   next
  1723     case (insert b F)
  1812     case False
  1724     show ?case
  1813     then have "{a \<in> A. f a = b} \<noteq> {}" by force
  1725     proof cases
  1814     with False show ?thesis by blast
  1726       assume "finite{a:A. f a = b}"
  1815   qed
  1727       hence "~ finite(A - {a:A. f a = b})" using \<open>\<not> finite A\<close> by simp
       
  1728       also have "A - {a:A. f a = b} = {a:A. f a \<noteq> b}" by blast
       
  1729       finally have "~ finite({a:A. f a \<noteq> b})" .
       
  1730       from insert(3)[OF _ this]
       
  1731       show ?thesis using insert(2,4) by simp (blast intro: rev_finite_subset)
       
  1732     next
       
  1733       assume 1: "~finite{a:A. f a = b}"
       
  1734       hence "{a \<in> A. f a = b} \<noteq> {}" by force
       
  1735       thus ?thesis using 1 by blast
       
  1736     qed
       
  1737   qed
       
  1738   from this[OF assms(2,1)] show ?thesis .
       
  1739 qed
  1816 qed
  1740 
  1817 
  1741 lemma pigeonhole_infinite_rel:
  1818 lemma pigeonhole_infinite_rel:
  1742 assumes "~finite A" and "finite B" and "ALL a:A. EX b:B. R a b"
  1819   assumes "\<not> finite A"
  1743 shows "EX b:B. ~finite{a:A. R a b}"
  1820     and "finite B"
  1744 proof -
  1821     and "\<forall>a\<in>A. \<exists>b\<in>B. R a b"
  1745    let ?F = "%a. {b:B. R a b}"
  1822   shows "\<exists>b\<in>B. \<not> finite {a:A. R a b}"
  1746    from finite_Pow_iff[THEN iffD2, OF \<open>finite B\<close>]
  1823 proof -
  1747    have "finite(?F ` A)" by(blast intro: rev_finite_subset)
  1824   let ?F = "\<lambda>a. {b\<in>B. R a b}"
  1748    from pigeonhole_infinite[where f = ?F, OF assms(1) this]
  1825   from finite_Pow_iff[THEN iffD2, OF \<open>finite B\<close>] have "finite (?F ` A)"
  1749    obtain a0 where "a0\<in>A" and 1: "\<not> finite {a\<in>A. ?F a = ?F a0}" ..
  1826     by (blast intro: rev_finite_subset)
  1750    obtain b0 where "b0 : B" and "R a0 b0" using \<open>a0:A\<close> assms(3) by blast
  1827   from pigeonhole_infinite [where f = ?F, OF assms(1) this]
  1751    { assume "finite{a:A. R a b0}"
  1828   obtain a0 where "a0 \<in> A" and 1: "\<not> finite {a\<in>A. ?F a = ?F a0}" ..
  1752      then have "finite {a\<in>A. ?F a = ?F a0}"
  1829   obtain b0 where "b0 \<in> B" and "R a0 b0"
  1753        using \<open>b0 : B\<close> \<open>R a0 b0\<close> by(blast intro: rev_finite_subset)
  1830     using \<open>a0 \<in> A\<close> assms(3) by blast
  1754    }
  1831   have "finite {a\<in>A. ?F a = ?F a0}" if "finite{a:A. R a b0}"
  1755    with 1 \<open>b0 : B\<close> show ?thesis by blast
  1832     using \<open>b0 \<in> B\<close> \<open>R a0 b0\<close> that by (blast intro: rev_finite_subset)
       
  1833   with 1 \<open>b0 : B\<close> show ?thesis
       
  1834     by blast
  1756 qed
  1835 qed
  1757 
  1836 
  1758 
  1837 
  1759 subsubsection \<open>Cardinality of sums\<close>
  1838 subsubsection \<open>Cardinality of sums\<close>
  1760 
  1839 
  1761 lemma card_Plus:
  1840 lemma card_Plus:
  1762   assumes "finite A" and "finite B"
  1841   assumes "finite A" "finite B"
  1763   shows "card (A <+> B) = card A + card B"
  1842   shows "card (A <+> B) = card A + card B"
  1764 proof -
  1843 proof -
  1765   have "Inl`A \<inter> Inr`B = {}" by fast
  1844   have "Inl`A \<inter> Inr`B = {}" by fast
  1766   with assms show ?thesis
  1845   with assms show ?thesis
  1767     unfolding Plus_def
  1846     by (simp add: Plus_def card_Un_disjoint card_image)
  1768     by (simp add: card_Un_disjoint card_image)
       
  1769 qed
  1847 qed
  1770 
  1848 
  1771 lemma card_Plus_conv_if:
  1849 lemma card_Plus_conv_if:
  1772   "card (A <+> B) = (if finite A \<and> finite B then card A + card B else 0)"
  1850   "card (A <+> B) = (if finite A \<and> finite B then card A + card B else 0)"
  1773   by (auto simp add: card_Plus)
  1851   by (auto simp add: card_Plus)
  1774 
  1852 
  1775 text \<open>Relates to equivalence classes.  Based on a theorem of F. Kamm\"uller.\<close>
  1853 text \<open>Relates to equivalence classes.  Based on a theorem of F. Kammüller.\<close>
  1776 
  1854 
  1777 lemma dvd_partition:
  1855 lemma dvd_partition:
  1778   assumes f: "finite (\<Union>C)" and "\<forall>c\<in>C. k dvd card c" "\<forall>c1\<in>C. \<forall>c2\<in>C. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {}"
  1856   assumes f: "finite (\<Union>C)"
  1779     shows "k dvd card (\<Union>C)"
  1857     and "\<forall>c\<in>C. k dvd card c" "\<forall>c1\<in>C. \<forall>c2\<in>C. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {}"
  1780 proof -
  1858   shows "k dvd card (\<Union>C)"
  1781   have "finite C" 
  1859 proof -
       
  1860   have "finite C"
  1782     by (rule finite_UnionD [OF f])
  1861     by (rule finite_UnionD [OF f])
  1783   then show ?thesis using assms
  1862   then show ?thesis
       
  1863     using assms
  1784   proof (induct rule: finite_induct)
  1864   proof (induct rule: finite_induct)
  1785     case empty show ?case by simp
  1865     case empty
       
  1866     show ?case by simp
  1786   next
  1867   next
  1787     case (insert c C)
  1868     case insert
  1788     then show ?case 
  1869     then show ?case
  1789       apply simp
  1870       apply simp
  1790       apply (subst card_Un_disjoint)
  1871       apply (subst card_Un_disjoint)
  1791       apply (auto simp add: disjoint_eq_subset_Compl)
  1872       apply (auto simp add: disjoint_eq_subset_Compl)
  1792       done
  1873       done
  1793   qed
  1874   qed
  1794 qed
  1875 qed
  1795 
  1876 
       
  1877 
  1796 subsubsection \<open>Relating injectivity and surjectivity\<close>
  1878 subsubsection \<open>Relating injectivity and surjectivity\<close>
  1797 
  1879 
  1798 lemma finite_surj_inj: assumes "finite A" "A \<subseteq> f ` A" shows "inj_on f A"
  1880 lemma finite_surj_inj:
  1799 proof -
  1881   assumes "finite A" "A \<subseteq> f ` A"
  1800   have "f ` A = A" 
  1882   shows "inj_on f A"
       
  1883 proof -
       
  1884   have "f ` A = A"
  1801     by (rule card_seteq [THEN sym]) (auto simp add: assms card_image_le)
  1885     by (rule card_seteq [THEN sym]) (auto simp add: assms card_image_le)
  1802   then show ?thesis using assms
  1886   then show ?thesis using assms
  1803     by (simp add: eq_card_imp_inj_on)
  1887     by (simp add: eq_card_imp_inj_on)
  1804 qed
  1888 qed
  1805 
  1889 
  1806 lemma finite_UNIV_surj_inj: fixes f :: "'a \<Rightarrow> 'a"
  1890 lemma finite_UNIV_surj_inj: "finite(UNIV:: 'a set) \<Longrightarrow> surj f \<Longrightarrow> inj f" for f :: "'a \<Rightarrow> 'a"
  1807 shows "finite(UNIV:: 'a set) \<Longrightarrow> surj f \<Longrightarrow> inj f"
  1891   by (blast intro: finite_surj_inj subset_UNIV)
  1808 by (blast intro: finite_surj_inj subset_UNIV)
  1892 
  1809 
  1893 lemma finite_UNIV_inj_surj: "finite(UNIV:: 'a set) \<Longrightarrow> inj f \<Longrightarrow> surj f" for f :: "'a \<Rightarrow> 'a"
  1810 lemma finite_UNIV_inj_surj: fixes f :: "'a \<Rightarrow> 'a"
  1894   by (fastforce simp:surj_def dest!: endo_inj_surj)
  1811 shows "finite(UNIV:: 'a set) \<Longrightarrow> inj f \<Longrightarrow> surj f"
  1895 
  1812 by(fastforce simp:surj_def dest!: endo_inj_surj)
  1896 corollary infinite_UNIV_nat [iff]: "\<not> finite (UNIV :: nat set)"
  1813 
       
  1814 corollary infinite_UNIV_nat [iff]:
       
  1815   "\<not> finite (UNIV :: nat set)"
       
  1816 proof
  1897 proof
  1817   assume "finite (UNIV :: nat set)"
  1898   assume "finite (UNIV :: nat set)"
  1818   with finite_UNIV_inj_surj [of Suc]
  1899   with finite_UNIV_inj_surj [of Suc] show False
  1819   show False by simp (blast dest: Suc_neq_Zero surjD)
  1900     by simp (blast dest: Suc_neq_Zero surjD)
  1820 qed
  1901 qed
  1821 
  1902 
  1822 lemma infinite_UNIV_char_0:
  1903 lemma infinite_UNIV_char_0: "\<not> finite (UNIV :: 'a::semiring_char_0 set)"
  1823   "\<not> finite (UNIV :: 'a::semiring_char_0 set)"
       
  1824 proof
  1904 proof
  1825   assume "finite (UNIV :: 'a set)"
  1905   assume "finite (UNIV :: 'a set)"
  1826   with subset_UNIV have "finite (range of_nat :: 'a set)"
  1906   with subset_UNIV have "finite (range of_nat :: 'a set)"
  1827     by (rule finite_subset)
  1907     by (rule finite_subset)
  1828   moreover have "inj (of_nat :: nat \<Rightarrow> 'a)"
  1908   moreover have "inj (of_nat :: nat \<Rightarrow> 'a)"
  1834 qed
  1914 qed
  1835 
  1915 
  1836 hide_const (open) Finite_Set.fold
  1916 hide_const (open) Finite_Set.fold
  1837 
  1917 
  1838 
  1918 
  1839 subsection "Infinite Sets"
  1919 subsection \<open>Infinite Sets\<close>
  1840 
  1920 
  1841 text \<open>
  1921 text \<open>
  1842   Some elementary facts about infinite sets, mostly by Stephan Merz.
  1922   Some elementary facts about infinite sets, mostly by Stephan Merz.
  1843   Beware! Because "infinite" merely abbreviates a negation, these
  1923   Beware! Because "infinite" merely abbreviates a negation, these
  1844   lemmas may not work well with \<open>blast\<close>.
  1924   lemmas may not work well with \<open>blast\<close>.
  1857 
  1937 
  1858 lemma infinite_remove: "infinite S \<Longrightarrow> infinite (S - {a})"
  1938 lemma infinite_remove: "infinite S \<Longrightarrow> infinite (S - {a})"
  1859   by simp
  1939   by simp
  1860 
  1940 
  1861 lemma Diff_infinite_finite:
  1941 lemma Diff_infinite_finite:
  1862   assumes T: "finite T" and S: "infinite S"
  1942   assumes "finite T" "infinite S"
  1863   shows "infinite (S - T)"
  1943   shows "infinite (S - T)"
  1864   using T
  1944   using \<open>finite T\<close>
  1865 proof induct
  1945 proof induct
  1866   from S
  1946   from \<open>infinite S\<close> show "infinite (S - {})"
  1867   show "infinite (S - {})" by auto
  1947     by auto
  1868 next
  1948 next
  1869   fix T x
  1949   fix T x
  1870   assume ih: "infinite (S - T)"
  1950   assume ih: "infinite (S - T)"
  1871   have "S - (insert x T) = (S - T) - {x}"
  1951   have "S - (insert x T) = (S - T) - {x}"
  1872     by (rule Diff_insert)
  1952     by (rule Diff_insert)
  1873   with ih
  1953   with ih show "infinite (S - (insert x T))"
  1874   show "infinite (S - (insert x T))"
       
  1875     by (simp add: infinite_remove)
  1954     by (simp add: infinite_remove)
  1876 qed
  1955 qed
  1877 
  1956 
  1878 lemma Un_infinite: "infinite S \<Longrightarrow> infinite (S \<union> T)"
  1957 lemma Un_infinite: "infinite S \<Longrightarrow> infinite (S \<union> T)"
  1879   by simp
  1958   by simp
  1880 
  1959 
  1881 lemma infinite_Un: "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"
  1960 lemma infinite_Un: "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"
  1882   by simp
  1961   by simp
  1883 
  1962 
  1884 lemma infinite_super:
  1963 lemma infinite_super:
  1885   assumes T: "S \<subseteq> T" and S: "infinite S"
  1964   assumes "S \<subseteq> T"
       
  1965     and "infinite S"
  1886   shows "infinite T"
  1966   shows "infinite T"
  1887 proof
  1967 proof
  1888   assume "finite T"
  1968   assume "finite T"
  1889   with T have "finite S" by (simp add: finite_subset)
  1969   with \<open>S \<subseteq> T\<close> have "finite S" by (simp add: finite_subset)
  1890   with S show False by simp
  1970   with \<open>infinite S\<close> show False by simp
  1891 qed
  1971 qed
  1892 
  1972 
  1893 proposition infinite_coinduct [consumes 1, case_names infinite]:
  1973 proposition infinite_coinduct [consumes 1, case_names infinite]:
  1894   assumes "X A"
  1974   assumes "X A"
  1895   and step: "\<And>A. X A \<Longrightarrow> \<exists>x\<in>A. X (A - {x}) \<or> infinite (A - {x})"
  1975     and step: "\<And>A. X A \<Longrightarrow> \<exists>x\<in>A. X (A - {x}) \<or> infinite (A - {x})"
  1896   shows "infinite A"
  1976   shows "infinite A"
  1897 proof
  1977 proof
  1898   assume "finite A"
  1978   assume "finite A"
  1899   then show False using \<open>X A\<close>
  1979   then show False
       
  1980     using \<open>X A\<close>
  1900   proof (induction rule: finite_psubset_induct)
  1981   proof (induction rule: finite_psubset_induct)
  1901     case (psubset A)
  1982     case (psubset A)
  1902     then obtain x where "x \<in> A" "X (A - {x}) \<or> infinite (A - {x})"
  1983     then obtain x where "x \<in> A" "X (A - {x}) \<or> infinite (A - {x})"
  1903       using local.step psubset.prems by blast
  1984       using local.step psubset.prems by blast
  1904     then have "X (A - {x})"
  1985     then have "X (A - {x})"
  1905       using psubset.hyps by blast
  1986       using psubset.hyps by blast
  1906     show False
  1987     show False
  1907       apply (rule psubset.IH [where B = "A - {x}"])
  1988       apply (rule psubset.IH [where B = "A - {x}"])
  1908       using \<open>x \<in> A\<close> apply blast
  1989       using \<open>x \<in> A\<close> apply blast
  1909       by (simp add: \<open>X (A - {x})\<close>)
  1990       apply (simp add: \<open>X (A - {x})\<close>)
       
  1991       done
  1910   qed
  1992   qed
  1911 qed
  1993 qed
  1912 
  1994 
  1913 text \<open>
  1995 text \<open>
  1914   For any function with infinite domain and finite range there is some
  1996   For any function with infinite domain and finite range there is some
  1916   particular, any infinite sequence of elements from a finite set
  1998   particular, any infinite sequence of elements from a finite set
  1917   contains some element that occurs infinitely often.
  1999   contains some element that occurs infinitely often.
  1918 \<close>
  2000 \<close>
  1919 
  2001 
  1920 lemma inf_img_fin_dom':
  2002 lemma inf_img_fin_dom':
  1921   assumes img: "finite (f ` A)" and dom: "infinite A"
  2003   assumes img: "finite (f ` A)"
       
  2004     and dom: "infinite A"
  1922   shows "\<exists>y \<in> f ` A. infinite (f -` {y} \<inter> A)"
  2005   shows "\<exists>y \<in> f ` A. infinite (f -` {y} \<inter> A)"
  1923 proof (rule ccontr)
  2006 proof (rule ccontr)
  1924   have "A \<subseteq> (\<Union>y\<in>f ` A. f -` {y} \<inter> A)" by auto
  2007   have "A \<subseteq> (\<Union>y\<in>f ` A. f -` {y} \<inter> A)" by auto
  1925   moreover
  2008   moreover assume "\<not> ?thesis"
  1926   assume "\<not> ?thesis"
       
  1927   with img have "finite (\<Union>y\<in>f ` A. f -` {y} \<inter> A)" by blast
  2009   with img have "finite (\<Union>y\<in>f ` A. f -` {y} \<inter> A)" by blast
  1928   ultimately have "finite A" by(rule finite_subset)
  2010   ultimately have "finite A" by (rule finite_subset)
  1929   with dom show False by contradiction
  2011   with dom show False by contradiction
  1930 qed
  2012 qed
  1931 
  2013 
  1932 lemma inf_img_fin_domE':
  2014 lemma inf_img_fin_domE':
  1933   assumes "finite (f ` A)" and "infinite A"
  2015   assumes "finite (f ` A)" and "infinite A"
  1935   using assms by (blast dest: inf_img_fin_dom')
  2017   using assms by (blast dest: inf_img_fin_dom')
  1936 
  2018 
  1937 lemma inf_img_fin_dom:
  2019 lemma inf_img_fin_dom:
  1938   assumes img: "finite (f`A)" and dom: "infinite A"
  2020   assumes img: "finite (f`A)" and dom: "infinite A"
  1939   shows "\<exists>y \<in> f`A. infinite (f -` {y})"
  2021   shows "\<exists>y \<in> f`A. infinite (f -` {y})"
  1940 using inf_img_fin_dom'[OF assms] by auto
  2022   using inf_img_fin_dom'[OF assms] by auto
  1941 
  2023 
  1942 lemma inf_img_fin_domE:
  2024 lemma inf_img_fin_domE:
  1943   assumes "finite (f`A)" and "infinite A"
  2025   assumes "finite (f`A)" and "infinite A"
  1944   obtains y where "y \<in> f`A" and "infinite (f -` {y})"
  2026   obtains y where "y \<in> f`A" and "infinite (f -` {y})"
  1945   using assms by (blast dest: inf_img_fin_dom)
  2027   using assms by (blast dest: inf_img_fin_dom)
  1946 
  2028 
  1947 proposition finite_image_absD:
  2029 proposition finite_image_absD: "finite (abs ` S) \<Longrightarrow> finite S"
  1948     fixes S :: "'a::linordered_ring set"
  2030   for S :: "'a::linordered_ring set"
  1949     shows "finite (abs ` S) \<Longrightarrow> finite S"
       
  1950   by (rule ccontr) (auto simp: abs_eq_iff vimage_def dest: inf_img_fin_dom)
  2031   by (rule ccontr) (auto simp: abs_eq_iff vimage_def dest: inf_img_fin_dom)
  1951 
  2032 
  1952 end
  2033 end