src/HOL/Finite_Set.thy
author haftmann
Sat, 02 Jul 2016 08:41:05 +0200
changeset 63365 5340fb6633d0
parent 63099 af0e964aad7b
child 63404 a95e7432d86c
permissions -rw-r--r--
more theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Finite_Set.thy
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
     2
    Author:     Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
55020
96b05fd2aee4 dissolved 'Fun_More_FP' (a BNF dependency)
blanchet
parents: 54870
diff changeset
     3
                with contributions by Jeremy Avigad and Andrei Popescu
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
     4
*)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
     5
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
     6
section \<open>Finite sets\<close>
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15124
diff changeset
     8
theory Finite_Set
62481
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62390
diff changeset
     9
imports Product_Type Sum_Type Fields
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15124
diff changeset
    10
begin
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    11
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
    12
subsection \<open>Predicate for finite sets\<close>
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    13
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61605
diff changeset
    14
context
62093
bd73a2279fcd more uniform treatment of package internals;
wenzelm
parents: 61890
diff changeset
    15
  notes [[inductive_internals]]
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61605
diff changeset
    16
begin
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61605
diff changeset
    17
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    18
inductive finite :: "'a set \<Rightarrow> bool"
22262
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 21733
diff changeset
    19
  where
96ba62dff413 Adapted to new inductive definition package.
berghofe
parents: 21733
diff changeset
    20
    emptyI [simp, intro!]: "finite {}"
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    21
  | insertI [simp, intro!]: "finite A \<Longrightarrow> finite (insert a A)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    22
61681
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61605
diff changeset
    23
end
ca53150406c9 option "inductive_defs" controls exposure of def and mono facts;
wenzelm
parents: 61605
diff changeset
    24
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
    25
simproc_setup finite_Collect ("finite (Collect P)") = \<open>K Set_Comprehension_Pointfree.simproc\<close>
48109
0a58f7eefba2 Integrated set comprehension pointfree simproc.
Rafal Kolanski <rafal.kolanski@nicta.com.au>
parents: 48063
diff changeset
    26
54611
31afce809794 set_comprehension_pointfree simproc causes to many surprises if enabled by default
traytel
parents: 54570
diff changeset
    27
declare [[simproc del: finite_Collect]]
31afce809794 set_comprehension_pointfree simproc causes to many surprises if enabled by default
traytel
parents: 54570
diff changeset
    28
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    29
lemma finite_induct [case_names empty insert, induct set: finite]:
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61778
diff changeset
    30
  \<comment> \<open>Discharging \<open>x \<notin> F\<close> entails extra work.\<close>
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    31
  assumes "finite F"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    32
  assumes "P {}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    33
    and insert: "\<And>x F. finite F \<Longrightarrow> x \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert x F)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    34
  shows "P F"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
    35
using \<open>finite F\<close>
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
    36
proof induct
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    37
  show "P {}" by fact
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    38
  fix x F assume F: "finite F" and P: "P F"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    39
  show "P (insert x F)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    40
  proof cases
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    41
    assume "x \<in> F"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    42
    hence "insert x F = F" by (rule insert_absorb)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    43
    with P show ?thesis by (simp only:)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    44
  next
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    45
    assume "x \<notin> F"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    46
    from F this P show ?thesis by (rule insert)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    47
  qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    48
qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    49
51622
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    50
lemma infinite_finite_induct [case_names infinite empty insert]:
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    51
  assumes infinite: "\<And>A. \<not> finite A \<Longrightarrow> P A"
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    52
  assumes empty: "P {}"
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    53
  assumes insert: "\<And>x F. finite F \<Longrightarrow> x \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert x F)"
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    54
  shows "P A"
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    55
proof (cases "finite A")
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    56
  case False with infinite show ?thesis .
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    57
next
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    58
  case True then show ?thesis by (induct A) (fact empty insert)+
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    59
qed
183f30bc11de convenient induction rule
haftmann
parents: 51598
diff changeset
    60
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    61
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
    62
subsubsection \<open>Choice principles\<close>
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    63
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61778
diff changeset
    64
lemma ex_new_if_finite: \<comment> "does not depend on def of finite at all"
14661
9ead82084de8 tuned notation;
wenzelm
parents: 14565
diff changeset
    65
  assumes "\<not> finite (UNIV :: 'a set)" and "finite A"
9ead82084de8 tuned notation;
wenzelm
parents: 14565
diff changeset
    66
  shows "\<exists>a::'a. a \<notin> A"
9ead82084de8 tuned notation;
wenzelm
parents: 14565
diff changeset
    67
proof -
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 27981
diff changeset
    68
  from assms have "A \<noteq> UNIV" by blast
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    69
  then show ?thesis by blast
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    70
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    71
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
    72
text \<open>A finite choice principle. Does not need the SOME choice operator.\<close>
15484
2636ec211ec8 fold and fol1 changes
nipkow
parents: 15483
diff changeset
    73
29923
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    74
lemma finite_set_choice:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    75
  "finite A \<Longrightarrow> \<forall>x\<in>A. \<exists>y. P x y \<Longrightarrow> \<exists>f. \<forall>x\<in>A. P x (f x)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    76
proof (induct rule: finite_induct)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    77
  case empty then show ?case by simp
29923
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    78
next
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    79
  case (insert a A)
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    80
  then obtain f b where f: "ALL x:A. P x (f x)" and ab: "P a b" by auto
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    81
  show ?case (is "EX f. ?P f")
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    82
  proof
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    83
    show "?P(%x. if x = a then b else f x)" using f ab by auto
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    84
  qed
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    85
qed
24f56736c56f added finite_set_choice
nipkow
parents: 29920
diff changeset
    86
23878
bd651ecd4b8a simplified HOL bootstrap
haftmann
parents: 23736
diff changeset
    87
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
    88
subsubsection \<open>Finite sets are the images of initial segments of natural numbers\<close>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
    89
15510
9de204d7b699 new foldSet proofs
paulson
parents: 15509
diff changeset
    90
lemma finite_imp_nat_seg_image_inj_on:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    91
  assumes "finite A" 
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    92
  shows "\<exists>(n::nat) f. A = f ` {i. i < n} \<and> inj_on f {i. i < n}"
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
    93
using assms
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
    94
proof induct
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
    95
  case empty
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    96
  show ?case
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    97
  proof
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
    98
    show "\<exists>f. {} = f ` {i::nat. i < 0} \<and> inj_on f {i. i < 0}" by simp 
15510
9de204d7b699 new foldSet proofs
paulson
parents: 15509
diff changeset
    99
  qed
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   100
next
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   101
  case (insert a A)
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23277
diff changeset
   102
  have notinA: "a \<notin> A" by fact
15510
9de204d7b699 new foldSet proofs
paulson
parents: 15509
diff changeset
   103
  from insert.hyps obtain n f
9de204d7b699 new foldSet proofs
paulson
parents: 15509
diff changeset
   104
    where "A = f ` {i::nat. i < n}" "inj_on f {i. i < n}" by blast
9de204d7b699 new foldSet proofs
paulson
parents: 15509
diff changeset
   105
  hence "insert a A = f(n:=a) ` {i. i < Suc n}"
9de204d7b699 new foldSet proofs
paulson
parents: 15509
diff changeset
   106
        "inj_on (f(n:=a)) {i. i < Suc n}" using notinA
9de204d7b699 new foldSet proofs
paulson
parents: 15509
diff changeset
   107
    by (auto simp add: image_def Ball_def inj_on_def less_Suc_eq)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   108
  thus ?case by blast
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   109
qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   110
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   111
lemma nat_seg_image_imp_finite:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   112
  "A = f ` {i::nat. i < n} \<Longrightarrow> finite A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   113
proof (induct n arbitrary: A)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   114
  case 0 thus ?case by simp
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   115
next
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   116
  case (Suc n)
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   117
  let ?B = "f ` {i. i < n}"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   118
  have finB: "finite ?B" by(rule Suc.hyps[OF refl])
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   119
  show ?case
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   120
  proof cases
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   121
    assume "\<exists>k<n. f n = f k"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   122
    hence "A = ?B" using Suc.prems by(auto simp:less_Suc_eq)
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   123
    thus ?thesis using finB by simp
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   124
  next
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   125
    assume "\<not>(\<exists> k<n. f n = f k)"
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   126
    hence "A = insert (f n) ?B" using Suc.prems by(auto simp:less_Suc_eq)
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   127
    thus ?thesis using finB by simp
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   128
  qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   129
qed
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   130
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   131
lemma finite_conv_nat_seg_image:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   132
  "finite A \<longleftrightarrow> (\<exists>(n::nat) f. A = f ` {i::nat. i < n})"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   133
  by (blast intro: nat_seg_image_imp_finite dest: finite_imp_nat_seg_image_inj_on)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   134
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 32705
diff changeset
   135
lemma finite_imp_inj_to_nat_seg:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   136
  assumes "finite A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   137
  shows "\<exists>f n::nat. f ` A = {i. i < n} \<and> inj_on f A"
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 32705
diff changeset
   138
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   139
  from finite_imp_nat_seg_image_inj_on[OF \<open>finite A\<close>]
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 32705
diff changeset
   140
  obtain f and n::nat where bij: "bij_betw f {i. i<n} A"
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 32705
diff changeset
   141
    by (auto simp:bij_betw_def)
33057
764547b68538 inv_onto -> inv_into
nipkow
parents: 32989
diff changeset
   142
  let ?f = "the_inv_into {i. i<n} f"
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 32705
diff changeset
   143
  have "inj_on ?f A & ?f ` A = {i. i<n}"
33057
764547b68538 inv_onto -> inv_into
nipkow
parents: 32989
diff changeset
   144
    by (fold bij_betw_def) (rule bij_betw_the_inv_into[OF bij])
32988
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 32705
diff changeset
   145
  thus ?thesis by blast
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 32705
diff changeset
   146
qed
d1d4d7a08a66 Inv -> inv_onto, inv abbr. inv_onto UNIV.
nipkow
parents: 32705
diff changeset
   147
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   148
lemma finite_Collect_less_nat [iff]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   149
  "finite {n::nat. n < k}"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44835
diff changeset
   150
  by (fastforce simp: finite_conv_nat_seg_image)
29920
b95f5b8b93dd more finiteness
nipkow
parents: 29918
diff changeset
   151
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   152
lemma finite_Collect_le_nat [iff]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   153
  "finite {n::nat. n \<le> k}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   154
  by (simp add: le_eq_less_or_eq Collect_disj_eq)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   155
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   156
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   157
subsubsection \<open>Finiteness and common set operations\<close>
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   158
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   159
lemma rev_finite_subset:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   160
  "finite B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> finite A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   161
proof (induct arbitrary: A rule: finite_induct)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   162
  case empty
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   163
  then show ?case by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   164
next
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   165
  case (insert x F A)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   166
  have A: "A \<subseteq> insert x F" and r: "A - {x} \<subseteq> F \<Longrightarrow> finite (A - {x})" by fact+
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   167
  show "finite A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   168
  proof cases
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   169
    assume x: "x \<in> A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   170
    with A have "A - {x} \<subseteq> F" by (simp add: subset_insert_iff)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   171
    with r have "finite (A - {x})" .
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   172
    hence "finite (insert x (A - {x}))" ..
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   173
    also have "insert x (A - {x}) = A" using x by (rule insert_Diff)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   174
    finally show ?thesis .
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   175
  next
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60585
diff changeset
   176
    show ?thesis when "A \<subseteq> F"
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 60585
diff changeset
   177
      using that by fact
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   178
    assume "x \<notin> A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   179
    with A show "A \<subseteq> F" by (simp add: subset_insert_iff)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   180
  qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   181
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   182
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   183
lemma finite_subset:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   184
  "A \<subseteq> B \<Longrightarrow> finite B \<Longrightarrow> finite A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   185
  by (rule rev_finite_subset)
29901
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29879
diff changeset
   186
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   187
lemma finite_UnI:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   188
  assumes "finite F" and "finite G"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   189
  shows "finite (F \<union> G)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   190
  using assms by induct simp_all
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
   191
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   192
lemma finite_Un [iff]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   193
  "finite (F \<union> G) \<longleftrightarrow> finite F \<and> finite G"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   194
  by (blast intro: finite_UnI finite_subset [of _ "F \<union> G"])
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
   195
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   196
lemma finite_insert [simp]: "finite (insert a A) \<longleftrightarrow> finite A"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   197
proof -
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   198
  have "finite {a} \<and> finite A \<longleftrightarrow> finite A" by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   199
  then have "finite ({a} \<union> A) \<longleftrightarrow> finite A" by (simp only: finite_Un)
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23277
diff changeset
   200
  then show ?thesis by simp
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   201
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   202
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   203
lemma finite_Int [simp, intro]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   204
  "finite F \<or> finite G \<Longrightarrow> finite (F \<inter> G)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   205
  by (blast intro: finite_subset)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   206
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   207
lemma finite_Collect_conjI [simp, intro]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   208
  "finite {x. P x} \<or> finite {x. Q x} \<Longrightarrow> finite {x. P x \<and> Q x}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   209
  by (simp add: Collect_conj_eq)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   210
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   211
lemma finite_Collect_disjI [simp]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   212
  "finite {x. P x \<or> Q x} \<longleftrightarrow> finite {x. P x} \<and> finite {x. Q x}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   213
  by (simp add: Collect_disj_eq)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   214
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   215
lemma finite_Diff [simp, intro]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   216
  "finite A \<Longrightarrow> finite (A - B)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   217
  by (rule finite_subset, rule Diff_subset)
29901
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29879
diff changeset
   218
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29879
diff changeset
   219
lemma finite_Diff2 [simp]:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   220
  assumes "finite B"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   221
  shows "finite (A - B) \<longleftrightarrow> finite A"
29901
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29879
diff changeset
   222
proof -
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   223
  have "finite A \<longleftrightarrow> finite((A - B) \<union> (A \<inter> B))" by (simp add: Un_Diff_Int)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   224
  also have "\<dots> \<longleftrightarrow> finite (A - B)" using \<open>finite B\<close> by simp
29901
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29879
diff changeset
   225
  finally show ?thesis ..
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29879
diff changeset
   226
qed
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29879
diff changeset
   227
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   228
lemma finite_Diff_insert [iff]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   229
  "finite (A - insert a B) \<longleftrightarrow> finite (A - B)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   230
proof -
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   231
  have "finite (A - B) \<longleftrightarrow> finite (A - B - {a})" by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   232
  moreover have "A - insert a B = A - B - {a}" by auto
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   233
  ultimately show ?thesis by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   234
qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   235
29901
f4b3f8fbf599 finiteness lemmas
nipkow
parents: 29879
diff changeset
   236
lemma finite_compl[simp]:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   237
  "finite (A :: 'a set) \<Longrightarrow> finite (- A) \<longleftrightarrow> finite (UNIV :: 'a set)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   238
  by (simp add: Compl_eq_Diff_UNIV)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   239
29916
f24137b42d9b more finiteness
nipkow
parents: 29903
diff changeset
   240
lemma finite_Collect_not[simp]:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   241
  "finite {x :: 'a. P x} \<Longrightarrow> finite {x. \<not> P x} \<longleftrightarrow> finite (UNIV :: 'a set)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   242
  by (simp add: Collect_neg_eq)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   243
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   244
lemma finite_Union [simp, intro]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   245
  "finite A \<Longrightarrow> (\<And>M. M \<in> A \<Longrightarrow> finite M) \<Longrightarrow> finite(\<Union>A)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   246
  by (induct rule: finite_induct) simp_all
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   247
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   248
lemma finite_UN_I [intro]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   249
  "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> finite (B a)) \<Longrightarrow> finite (\<Union>a\<in>A. B a)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   250
  by (induct rule: finite_induct) simp_all
29903
2c0046b26f80 more finiteness changes
nipkow
parents: 29901
diff changeset
   251
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   252
lemma finite_UN [simp]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   253
  "finite A \<Longrightarrow> finite (UNION A B) \<longleftrightarrow> (\<forall>x\<in>A. finite (B x))"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   254
  by (blast intro: finite_subset)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   255
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   256
lemma finite_Inter [intro]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   257
  "\<exists>A\<in>M. finite A \<Longrightarrow> finite (\<Inter>M)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   258
  by (blast intro: Inter_lower finite_subset)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   259
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   260
lemma finite_INT [intro]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   261
  "\<exists>x\<in>I. finite (A x) \<Longrightarrow> finite (\<Inter>x\<in>I. A x)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   262
  by (blast intro: INT_lower finite_subset)
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   263
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   264
lemma finite_imageI [simp, intro]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   265
  "finite F \<Longrightarrow> finite (h ` F)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   266
  by (induct rule: finite_induct) simp_all
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   267
31768
159cd6b5e5d4 lemma finite_image_set by Jeremy Avigad
haftmann
parents: 31465
diff changeset
   268
lemma finite_image_set [simp]:
159cd6b5e5d4 lemma finite_image_set by Jeremy Avigad
haftmann
parents: 31465
diff changeset
   269
  "finite {x. P x} \<Longrightarrow> finite { f x | x. P x }"
159cd6b5e5d4 lemma finite_image_set by Jeremy Avigad
haftmann
parents: 31465
diff changeset
   270
  by (simp add: image_Collect [symmetric])
159cd6b5e5d4 lemma finite_image_set by Jeremy Avigad
haftmann
parents: 31465
diff changeset
   271
59504
8c6747dba731 New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents: 59336
diff changeset
   272
lemma finite_image_set2:
8c6747dba731 New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents: 59336
diff changeset
   273
  "finite {x. P x} \<Longrightarrow> finite {y. Q y} \<Longrightarrow> finite {f x y | x y. P x \<and> Q y}"
8c6747dba731 New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents: 59336
diff changeset
   274
  by (rule finite_subset [where B = "\<Union>x \<in> {x. P x}. \<Union>y \<in> {y. Q y}. {f x y}"]) auto
8c6747dba731 New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents: 59336
diff changeset
   275
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   276
lemma finite_imageD:
42206
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   277
  assumes "finite (f ` A)" and "inj_on f A"
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   278
  shows "finite A"
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   279
using assms
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   280
proof (induct "f ` A" arbitrary: A)
42206
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   281
  case empty then show ?case by simp
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   282
next
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   283
  case (insert x B)
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   284
  then have B_A: "insert x B = f ` A" by simp
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   285
  then obtain y where "x = f y" and "y \<in> A" by blast
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   286
  from B_A \<open>x \<notin> B\<close> have "B = f ` A - {x}" by blast
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   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})" 
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 59520
diff changeset
   288
    by (simp add: inj_on_image_set_diff Set.Diff_subset)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   289
  moreover from \<open>inj_on f A\<close> have "inj_on f (A - {y})" by (rule inj_on_diff)
42206
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   290
  ultimately have "finite (A - {y})" by (rule insert.hyps)
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   291
  then show "finite A" by simp
0920f709610f tuned proof
haftmann
parents: 41988
diff changeset
   292
qed
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   293
62618
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62481
diff changeset
   294
lemma finite_image_iff:
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62481
diff changeset
   295
  assumes "inj_on f A"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62481
diff changeset
   296
  shows "finite (f ` A) \<longleftrightarrow> finite A"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62481
diff changeset
   297
using assms finite_imageD by blast
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62481
diff changeset
   298
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   299
lemma finite_surj:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   300
  "finite A \<Longrightarrow> B \<subseteq> f ` A \<Longrightarrow> finite B"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   301
  by (erule finite_subset) (rule finite_imageI)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   302
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   303
lemma finite_range_imageI:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   304
  "finite (range g) \<Longrightarrow> finite (range (\<lambda>x. f (g x)))"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   305
  by (drule finite_imageI) (simp add: range_composition)
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   306
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   307
lemma finite_subset_image:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   308
  assumes "finite B"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   309
  shows "B \<subseteq> f ` A \<Longrightarrow> \<exists>C\<subseteq>A. finite C \<and> B = f ` C"
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   310
using assms
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   311
proof induct
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   312
  case empty then show ?case by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   313
next
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   314
  case insert then show ?case
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   315
    by (clarsimp simp del: image_insert simp add: image_insert [symmetric])
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   316
       blast
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   317
qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   318
43991
f4a7697011c5 finite vimage on arbitrary domains
hoelzl
parents: 43866
diff changeset
   319
lemma finite_vimage_IntI:
f4a7697011c5 finite vimage on arbitrary domains
hoelzl
parents: 43866
diff changeset
   320
  "finite F \<Longrightarrow> inj_on h A \<Longrightarrow> finite (h -` F \<inter> A)"
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   321
  apply (induct rule: finite_induct)
21575
89463ae2612d tuned proofs;
wenzelm
parents: 21409
diff changeset
   322
   apply simp_all
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   323
  apply (subst vimage_insert)
43991
f4a7697011c5 finite vimage on arbitrary domains
hoelzl
parents: 43866
diff changeset
   324
  apply (simp add: finite_subset [OF inj_on_vimage_singleton] Int_Un_distrib2)
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   325
  done
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   326
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   327
lemma finite_finite_vimage_IntI:
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   328
  assumes "finite F" and "\<And>y. y \<in> F \<Longrightarrow> finite ((h -` {y}) \<inter> A)"
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   329
  shows "finite (h -` F \<inter> A)"
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   330
proof -
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   331
  have *: "h -` F \<inter> A = (\<Union> y\<in>F. (h -` {y}) \<inter> A)"
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   332
    by blast
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   333
  show ?thesis
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   334
    by (simp only: * assms finite_UN_I)
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   335
qed
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61681
diff changeset
   336
43991
f4a7697011c5 finite vimage on arbitrary domains
hoelzl
parents: 43866
diff changeset
   337
lemma finite_vimageI:
f4a7697011c5 finite vimage on arbitrary domains
hoelzl
parents: 43866
diff changeset
   338
  "finite F \<Longrightarrow> inj h \<Longrightarrow> finite (h -` F)"
f4a7697011c5 finite vimage on arbitrary domains
hoelzl
parents: 43866
diff changeset
   339
  using finite_vimage_IntI[of F h UNIV] by auto
f4a7697011c5 finite vimage on arbitrary domains
hoelzl
parents: 43866
diff changeset
   340
59519
2fb0c0fc62a3 add more general version of finite_vimageD
Andreas Lochbihler
parents: 59504
diff changeset
   341
lemma finite_vimageD': "\<lbrakk> finite (f -` A); A \<subseteq> range f \<rbrakk> \<Longrightarrow> finite A"
2fb0c0fc62a3 add more general version of finite_vimageD
Andreas Lochbihler
parents: 59504
diff changeset
   342
by(auto simp add: subset_image_iff intro: finite_subset[rotated])
2fb0c0fc62a3 add more general version of finite_vimageD
Andreas Lochbihler
parents: 59504
diff changeset
   343
2fb0c0fc62a3 add more general version of finite_vimageD
Andreas Lochbihler
parents: 59504
diff changeset
   344
lemma finite_vimageD: "\<lbrakk> finite (h -` F); surj h \<rbrakk> \<Longrightarrow> finite F"
2fb0c0fc62a3 add more general version of finite_vimageD
Andreas Lochbihler
parents: 59504
diff changeset
   345
by(auto dest: finite_vimageD')
34111
1b015caba46c add lemmas rev_finite_subset, finite_vimageD, finite_vimage_iff
huffman
parents: 34007
diff changeset
   346
1b015caba46c add lemmas rev_finite_subset, finite_vimageD, finite_vimage_iff
huffman
parents: 34007
diff changeset
   347
lemma finite_vimage_iff: "bij h \<Longrightarrow> finite (h -` F) \<longleftrightarrow> finite F"
1b015caba46c add lemmas rev_finite_subset, finite_vimageD, finite_vimage_iff
huffman
parents: 34007
diff changeset
   348
  unfolding bij_def by (auto elim: finite_vimageD finite_vimageI)
1b015caba46c add lemmas rev_finite_subset, finite_vimageD, finite_vimage_iff
huffman
parents: 34007
diff changeset
   349
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   350
lemma finite_Collect_bex [simp]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   351
  assumes "finite A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   352
  shows "finite {x. \<exists>y\<in>A. Q x y} \<longleftrightarrow> (\<forall>y\<in>A. finite {x. Q x y})"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   353
proof -
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   354
  have "{x. \<exists>y\<in>A. Q x y} = (\<Union>y\<in>A. {x. Q x y})" by auto
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   355
  with assms show ?thesis by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   356
qed
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   357
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   358
lemma finite_Collect_bounded_ex [simp]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   359
  assumes "finite {y. P y}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   360
  shows "finite {x. \<exists>y. P y \<and> Q x y} \<longleftrightarrow> (\<forall>y. P y \<longrightarrow> finite {x. Q x y})"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   361
proof -
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   362
  have "{x. EX y. P y & Q x y} = (\<Union>y\<in>{y. P y}. {x. Q x y})" by auto
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   363
  with assms show ?thesis by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   364
qed
29920
b95f5b8b93dd more finiteness
nipkow
parents: 29918
diff changeset
   365
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   366
lemma finite_Plus:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   367
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A <+> B)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   368
  by (simp add: Plus_def)
17022
b257300c3a9c added Brian Hufmann's finite instances
nipkow
parents: 16775
diff changeset
   369
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   370
lemma finite_PlusD: 
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   371
  fixes A :: "'a set" and B :: "'b set"
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   372
  assumes fin: "finite (A <+> B)"
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   373
  shows "finite A" "finite B"
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   374
proof -
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   375
  have "Inl ` A \<subseteq> A <+> B" by auto
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   376
  then have "finite (Inl ` A :: ('a + 'b) set)" using fin by (rule finite_subset)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   377
  then show "finite A" by (rule finite_imageD) (auto intro: inj_onI)
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   378
next
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   379
  have "Inr ` B \<subseteq> A <+> B" by auto
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   380
  then have "finite (Inr ` B :: ('a + 'b) set)" using fin by (rule finite_subset)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   381
  then show "finite B" by (rule finite_imageD) (auto intro: inj_onI)
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   382
qed
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   383
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   384
lemma finite_Plus_iff [simp]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   385
  "finite (A <+> B) \<longleftrightarrow> finite A \<and> finite B"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   386
  by (auto intro: finite_PlusD finite_Plus)
31080
21ffc770ebc0 lemmas by Andreas Lochbihler
nipkow
parents: 31017
diff changeset
   387
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   388
lemma finite_Plus_UNIV_iff [simp]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   389
  "finite (UNIV :: ('a + 'b) set) \<longleftrightarrow> finite (UNIV :: 'a set) \<and> finite (UNIV :: 'b set)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   390
  by (subst UNIV_Plus_UNIV [symmetric]) (rule finite_Plus_iff)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   391
40786
0a54cfc9add3 gave more standard finite set rules simp and intro attribute
nipkow
parents: 40716
diff changeset
   392
lemma finite_SigmaI [simp, intro]:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   393
  "finite A \<Longrightarrow> (\<And>a. a\<in>A \<Longrightarrow> finite (B a)) ==> finite (SIGMA a:A. B a)"
40786
0a54cfc9add3 gave more standard finite set rules simp and intro attribute
nipkow
parents: 40716
diff changeset
   394
  by (unfold Sigma_def) blast
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   395
51290
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   396
lemma finite_SigmaI2:
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   397
  assumes "finite {x\<in>A. B x \<noteq> {}}"
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   398
  and "\<And>a. a \<in> A \<Longrightarrow> finite (B a)"
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   399
  shows "finite (Sigma A B)"
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   400
proof -
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   401
  from assms have "finite (Sigma {x\<in>A. B x \<noteq> {}} B)" by auto
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   402
  also have "Sigma {x:A. B x \<noteq> {}} B = Sigma A B" by auto
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   403
  finally show ?thesis .
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   404
qed
c48477e76de5 added lemma
Andreas Lochbihler
parents: 49806
diff changeset
   405
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   406
lemma finite_cartesian_product:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   407
  "finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<times> B)"
15402
97204f3b4705 REorganized Finite_Set
nipkow
parents: 15392
diff changeset
   408
  by (rule finite_SigmaI)
97204f3b4705 REorganized Finite_Set
nipkow
parents: 15392
diff changeset
   409
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   410
lemma finite_Prod_UNIV:
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   411
  "finite (UNIV :: 'a set) \<Longrightarrow> finite (UNIV :: 'b set) \<Longrightarrow> finite (UNIV :: ('a \<times> 'b) set)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   412
  by (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   413
15409
a063687d24eb new and stronger lemmas and improved simplification for finite sets
paulson
parents: 15402
diff changeset
   414
lemma finite_cartesian_productD1:
42207
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   415
  assumes "finite (A \<times> B)" and "B \<noteq> {}"
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   416
  shows "finite A"
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   417
proof -
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   418
  from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   419
    by (auto simp add: finite_conv_nat_seg_image)
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   420
  then have "fst ` (A \<times> B) = fst ` f ` {i::nat. i < n}" by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   421
  with \<open>B \<noteq> {}\<close> have "A = (fst \<circ> f) ` {i::nat. i < n}"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55096
diff changeset
   422
    by (simp add: image_comp)
42207
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   423
  then have "\<exists>n f. A = f ` {i::nat. i < n}" by blast
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   424
  then show ?thesis
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   425
    by (auto simp add: finite_conv_nat_seg_image)
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   426
qed
15409
a063687d24eb new and stronger lemmas and improved simplification for finite sets
paulson
parents: 15402
diff changeset
   427
a063687d24eb new and stronger lemmas and improved simplification for finite sets
paulson
parents: 15402
diff changeset
   428
lemma finite_cartesian_productD2:
42207
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   429
  assumes "finite (A \<times> B)" and "A \<noteq> {}"
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   430
  shows "finite B"
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   431
proof -
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   432
  from assms obtain n f where "A \<times> B = f ` {i::nat. i < n}"
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   433
    by (auto simp add: finite_conv_nat_seg_image)
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   434
  then have "snd ` (A \<times> B) = snd ` f ` {i::nat. i < n}" by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   435
  with \<open>A \<noteq> {}\<close> have "B = (snd \<circ> f) ` {i::nat. i < n}"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55096
diff changeset
   436
    by (simp add: image_comp)
42207
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   437
  then have "\<exists>n f. B = f ` {i::nat. i < n}" by blast
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   438
  then show ?thesis
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   439
    by (auto simp add: finite_conv_nat_seg_image)
2bda5eddadf3 tuned proofs
haftmann
parents: 42206
diff changeset
   440
qed
15409
a063687d24eb new and stronger lemmas and improved simplification for finite sets
paulson
parents: 15402
diff changeset
   441
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56218
diff changeset
   442
lemma finite_cartesian_product_iff:
e7fd64f82876 add various lemmas
hoelzl
parents: 56218
diff changeset
   443
  "finite (A \<times> B) \<longleftrightarrow> (A = {} \<or> B = {} \<or> (finite A \<and> finite B))"
e7fd64f82876 add various lemmas
hoelzl
parents: 56218
diff changeset
   444
  by (auto dest: finite_cartesian_productD1 finite_cartesian_productD2 finite_cartesian_product)
e7fd64f82876 add various lemmas
hoelzl
parents: 56218
diff changeset
   445
48175
fea68365c975 add finiteness lemmas for 'a * 'b and 'a set
Andreas Lochbihler
parents: 48128
diff changeset
   446
lemma finite_prod: 
fea68365c975 add finiteness lemmas for 'a * 'b and 'a set
Andreas Lochbihler
parents: 48128
diff changeset
   447
  "finite (UNIV :: ('a \<times> 'b) set) \<longleftrightarrow> finite (UNIV :: 'a set) \<and> finite (UNIV :: 'b set)"
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56218
diff changeset
   448
  using finite_cartesian_product_iff[of UNIV UNIV] by simp
48175
fea68365c975 add finiteness lemmas for 'a * 'b and 'a set
Andreas Lochbihler
parents: 48128
diff changeset
   449
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   450
lemma finite_Pow_iff [iff]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   451
  "finite (Pow A) \<longleftrightarrow> finite A"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   452
proof
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   453
  assume "finite (Pow A)"
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   454
  then have "finite ((%x. {x}) ` A)" by (blast intro: finite_subset)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   455
  then show "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   456
next
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   457
  assume "finite A"
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   458
  then show "finite (Pow A)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35171
diff changeset
   459
    by induct (simp_all add: Pow_insert)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   460
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   461
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   462
corollary finite_Collect_subsets [simp, intro]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   463
  "finite A \<Longrightarrow> finite {B. B \<subseteq> A}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   464
  by (simp add: Pow_def [symmetric])
29918
214755b03df3 more finiteness
nipkow
parents: 29916
diff changeset
   465
48175
fea68365c975 add finiteness lemmas for 'a * 'b and 'a set
Andreas Lochbihler
parents: 48128
diff changeset
   466
lemma finite_set: "finite (UNIV :: 'a set set) \<longleftrightarrow> finite (UNIV :: 'a set)"
fea68365c975 add finiteness lemmas for 'a * 'b and 'a set
Andreas Lochbihler
parents: 48128
diff changeset
   467
by(simp only: finite_Pow_iff Pow_UNIV[symmetric])
fea68365c975 add finiteness lemmas for 'a * 'b and 'a set
Andreas Lochbihler
parents: 48128
diff changeset
   468
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   469
lemma finite_UnionD: "finite(\<Union>A) \<Longrightarrow> finite A"
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   470
  by (blast intro: finite_subset [OF subset_Pow_Union])
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   471
53820
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   472
lemma finite_set_of_finite_funs: assumes "finite A" "finite B"
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   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")
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   474
proof-
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   475
  let ?F = "\<lambda>f. {(a,b). a \<in> A \<and> b = f a}"
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   476
  have "?F ` ?S \<subseteq> Pow(A \<times> B)" by auto
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   477
  from finite_subset[OF this] assms have 1: "finite (?F ` ?S)" by simp
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   478
  have 2: "inj_on ?F ?S"
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   479
    by(fastforce simp add: inj_on_def set_eq_iff fun_eq_iff)
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   480
  show ?thesis by(rule finite_imageD[OF 1 2])
9c7e97d67b45 added lemmas
nipkow
parents: 53374
diff changeset
   481
qed
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   482
58195
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   483
lemma not_finite_existsD:
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   484
  assumes "\<not> finite {a. P a}"
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   485
  shows "\<exists>a. P a"
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   486
proof (rule classical)
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   487
  assume "\<not> (\<exists>a. P a)"
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   488
  with assms show ?thesis by auto
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   489
qed
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   490
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   491
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   492
subsubsection \<open>Further induction rules on finite sets\<close>
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   493
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   494
lemma finite_ne_induct [case_names singleton insert, consumes 2]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   495
  assumes "finite F" and "F \<noteq> {}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   496
  assumes "\<And>x. P {x}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   497
    and "\<And>x F. finite F \<Longrightarrow> F \<noteq> {} \<Longrightarrow> x \<notin> F \<Longrightarrow> P F  \<Longrightarrow> P (insert x F)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   498
  shows "P F"
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   499
using assms
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   500
proof induct
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   501
  case empty then show ?case by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   502
next
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   503
  case (insert x F) then show ?case by cases auto
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   504
qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   505
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   506
lemma finite_subset_induct [consumes 2, case_names empty insert]:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   507
  assumes "finite F" and "F \<subseteq> A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   508
  assumes empty: "P {}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   509
    and insert: "\<And>a F. finite F \<Longrightarrow> a \<in> A \<Longrightarrow> a \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert a F)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   510
  shows "P F"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   511
using \<open>finite F\<close> \<open>F \<subseteq> A\<close>
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   512
proof induct
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   513
  show "P {}" by fact
31441
428e4caf2299 finite lemmas
nipkow
parents: 31438
diff changeset
   514
next
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   515
  fix x F
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   516
  assume "finite F" and "x \<notin> F" and
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   517
    P: "F \<subseteq> A \<Longrightarrow> P F" and i: "insert x F \<subseteq> A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   518
  show "P (insert x F)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   519
  proof (rule insert)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   520
    from i show "x \<in> A" by blast
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   521
    from i have "F \<subseteq> A" by blast
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   522
    with P show "P F" .
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   523
    show "finite F" by fact
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   524
    show "x \<notin> F" by fact
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   525
  qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   526
qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   527
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   528
lemma finite_empty_induct:
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   529
  assumes "finite A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   530
  assumes "P A"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   531
    and remove: "\<And>a A. finite A \<Longrightarrow> a \<in> A \<Longrightarrow> P A \<Longrightarrow> P (A - {a})"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   532
  shows "P {}"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   533
proof -
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   534
  have "\<And>B. B \<subseteq> A \<Longrightarrow> P (A - B)"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   535
  proof -
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   536
    fix B :: "'a set"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   537
    assume "B \<subseteq> A"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   538
    with \<open>finite A\<close> have "finite B" by (rule rev_finite_subset)
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   539
    from this \<open>B \<subseteq> A\<close> show "P (A - B)"
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   540
    proof induct
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   541
      case empty
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   542
      from \<open>P A\<close> show ?case by simp
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   543
    next
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   544
      case (insert b B)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   545
      have "P (A - B - {b})"
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   546
      proof (rule remove)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   547
        from \<open>finite A\<close> show "finite (A - B)" by induct auto
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   548
        from insert show "b \<in> A - B" by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   549
        from insert show "P (A - B)" by simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   550
      qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   551
      also have "A - B - {b} = A - insert b B" by (rule Diff_insert [symmetric])
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   552
      finally show ?case .
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   553
    qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   554
  qed
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   555
  then have "P (A - A)" by blast
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   556
  then show ?thesis by simp
31441
428e4caf2299 finite lemmas
nipkow
parents: 31438
diff changeset
   557
qed
428e4caf2299 finite lemmas
nipkow
parents: 31438
diff changeset
   558
58195
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   559
lemma finite_update_induct [consumes 1, case_names const update]:
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   560
  assumes finite: "finite {a. f a \<noteq> c}"
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   561
  assumes const: "P (\<lambda>a. c)"
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   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))"
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   563
  shows "P f"
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   564
using finite proof (induct "{a. f a \<noteq> c}" arbitrary: f)
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   565
  case empty with const show ?case by simp
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   566
next
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   567
  case (insert a A)
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   568
  then have "A = {a'. (f(a := c)) a' \<noteq> c}" and "f a \<noteq> c"
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   569
    by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   570
  with \<open>finite A\<close> have "finite {a'. (f(a := c)) a' \<noteq> c}"
58195
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   571
    by simp
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   572
  have "(f(a := c)) a = c"
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   573
    by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   574
  from insert \<open>A = {a'. (f(a := c)) a' \<noteq> c}\<close> have "P (f(a := c))"
58195
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   575
    by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   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))"
58195
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   577
    by (rule update)
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   578
  then show ?case by simp
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   579
qed
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   580
1fee63e0377d added various facts
haftmann
parents: 57598
diff changeset
   581
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61778
diff changeset
   582
subsection \<open>Class \<open>finite\<close>\<close>
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   583
29797
08ef36ed2f8a handling type classes without parameters
haftmann
parents: 29675
diff changeset
   584
class finite =
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60762
diff changeset
   585
  assumes finite_UNIV: "finite (UNIV :: 'a set)"
27430
1e25ac05cd87 prove lemma finite in context of finite class
huffman
parents: 27418
diff changeset
   586
begin
1e25ac05cd87 prove lemma finite in context of finite class
huffman
parents: 27418
diff changeset
   587
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60762
diff changeset
   588
lemma finite [simp]: "finite (A :: 'a set)"
26441
7914697ff104 no "attach UNIV" any more
haftmann
parents: 26146
diff changeset
   589
  by (rule subset_UNIV finite_UNIV finite_subset)+
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   590
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60762
diff changeset
   591
lemma finite_code [code]: "finite (A :: 'a set) \<longleftrightarrow> True"
40922
4d0f96a54e76 adding code equation for finiteness of finite types
bulwahn
parents: 40786
diff changeset
   592
  by simp
4d0f96a54e76 adding code equation for finiteness of finite types
bulwahn
parents: 40786
diff changeset
   593
27430
1e25ac05cd87 prove lemma finite in context of finite class
huffman
parents: 27418
diff changeset
   594
end
1e25ac05cd87 prove lemma finite in context of finite class
huffman
parents: 27418
diff changeset
   595
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   596
instance prod :: (finite, finite) finite
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
   597
  by standard (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite)
26146
61cb176d0385 tuned proofs
haftmann
parents: 26041
diff changeset
   598
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   599
lemma inj_graph: "inj (%f. {(x, y). y = f x})"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   600
  by (rule inj_onI, auto simp add: set_eq_iff fun_eq_iff)
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   601
26146
61cb176d0385 tuned proofs
haftmann
parents: 26041
diff changeset
   602
instance "fun" :: (finite, finite) finite
61cb176d0385 tuned proofs
haftmann
parents: 26041
diff changeset
   603
proof
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   604
  show "finite (UNIV :: ('a => 'b) set)"
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   605
  proof (rule finite_imageD)
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   606
    let ?graph = "%f::'a => 'b. {(x, y). y = f x}"
26792
f2d75fd23124 - Deleted code setup for finite and card
berghofe
parents: 26757
diff changeset
   607
    have "range ?graph \<subseteq> Pow UNIV" by simp
f2d75fd23124 - Deleted code setup for finite and card
berghofe
parents: 26757
diff changeset
   608
    moreover have "finite (Pow (UNIV :: ('a * 'b) set))"
f2d75fd23124 - Deleted code setup for finite and card
berghofe
parents: 26757
diff changeset
   609
      by (simp only: finite_Pow_iff finite)
f2d75fd23124 - Deleted code setup for finite and card
berghofe
parents: 26757
diff changeset
   610
    ultimately show "finite (range ?graph)"
f2d75fd23124 - Deleted code setup for finite and card
berghofe
parents: 26757
diff changeset
   611
      by (rule finite_subset)
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   612
    show "inj ?graph" by (rule inj_graph)
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   613
  qed
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   614
qed
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   615
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   616
instance bool :: finite
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
   617
  by standard (simp add: UNIV_bool)
44831
haftmann
parents: 43991
diff changeset
   618
45962
fc77947a7db4 finite type class instance for `set`
haftmann
parents: 45166
diff changeset
   619
instance set :: (finite) finite
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
   620
  by standard (simp only: Pow_UNIV [symmetric] finite_Pow_iff finite)
45962
fc77947a7db4 finite type class instance for `set`
haftmann
parents: 45166
diff changeset
   621
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   622
instance unit :: finite
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
   623
  by standard (simp add: UNIV_unit)
44831
haftmann
parents: 43991
diff changeset
   624
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46146
diff changeset
   625
instance sum :: (finite, finite) finite
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
   626
  by standard (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite)
27981
feb0c01cf0fb tuned import order
haftmann
parents: 27611
diff changeset
   627
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   628
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   629
subsection \<open>A basic fold functional for finite sets\<close>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   630
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   631
text \<open>The intended behaviour is
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61778
diff changeset
   632
\<open>fold f z {x\<^sub>1, ..., x\<^sub>n} = f x\<^sub>1 (\<dots> (f x\<^sub>n z)\<dots>)\<close>
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61778
diff changeset
   633
if \<open>f\<close> is ``left-commutative'':
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   634
\<close>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   635
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   636
locale comp_fun_commute =
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   637
  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   638
  assumes comp_fun_commute: "f y \<circ> f x = f x \<circ> f y"
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   639
begin
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   640
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   641
lemma fun_left_comm: "f y (f x z) = f x (f y z)"
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   642
  using comp_fun_commute by (simp add: fun_eq_iff)
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   643
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   644
lemma commute_left_comp:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   645
  "f y \<circ> (f x \<circ> g) = f x \<circ> (f y \<circ> g)"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   646
  by (simp add: o_assoc comp_fun_commute)
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   647
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   648
end
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   649
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   650
inductive fold_graph :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   651
for f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" and z :: 'b where
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   652
  emptyI [intro]: "fold_graph f z {} z" |
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   653
  insertI [intro]: "x \<notin> A \<Longrightarrow> fold_graph f z A y
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   654
      \<Longrightarrow> fold_graph f z (insert x A) (f x y)"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   655
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   656
inductive_cases empty_fold_graphE [elim!]: "fold_graph f z {} x"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   657
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   658
definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b" where
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   659
  "fold f z A = (if finite A then (THE y. fold_graph f z A y) else z)"
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   660
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   661
text\<open>A tempting alternative for the definiens is
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   662
@{term "if finite A then THE y. fold_graph f z A y else e"}.
15498
3988e90613d4 comment
paulson
parents: 15497
diff changeset
   663
It allows the removal of finiteness assumptions from the theorems
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61778
diff changeset
   664
\<open>fold_comm\<close>, \<open>fold_reindex\<close> and \<open>fold_distrib\<close>.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   665
The proofs become ugly. It is not worth the effort. (???)\<close>
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   666
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   667
lemma finite_imp_fold_graph: "finite A \<Longrightarrow> \<exists>x. fold_graph f z A x"
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
   668
by (induct rule: finite_induct) auto
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   669
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   670
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   671
subsubsection\<open>From @{const fold_graph} to @{term fold}\<close>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   672
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   673
context comp_fun_commute
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   674
begin
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   675
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   676
lemma fold_graph_finite:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   677
  assumes "fold_graph f z A y"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   678
  shows "finite A"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   679
  using assms by induct simp_all
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   680
36045
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   681
lemma fold_graph_insertE_aux:
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   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'"
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   683
proof (induct set: fold_graph)
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   684
  case (insertI x A y) show ?case
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   685
  proof (cases "x = a")
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   686
    assume "x = a" with insertI show ?case by auto
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   687
  next
36045
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   688
    assume "x \<noteq> a"
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   689
    then obtain y' where y: "y = f a y'" and y': "fold_graph f z (A - {a}) y'"
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   690
      using insertI by auto
42875
d1aad0957eb2 tuned proofs
haftmann
parents: 42873
diff changeset
   691
    have "f x y = f a (f x y')"
36045
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   692
      unfolding y by (rule fun_left_comm)
42875
d1aad0957eb2 tuned proofs
haftmann
parents: 42873
diff changeset
   693
    moreover have "fold_graph f z (insert x A - {a}) (f x y')"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   694
      using y' and \<open>x \<noteq> a\<close> and \<open>x \<notin> A\<close>
36045
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   695
      by (simp add: insert_Diff_if fold_graph.insertI)
42875
d1aad0957eb2 tuned proofs
haftmann
parents: 42873
diff changeset
   696
    ultimately show ?case by fast
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   697
  qed
36045
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   698
qed simp
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   699
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   700
lemma fold_graph_insertE:
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   701
  assumes "fold_graph f z (insert x A) v" and "x \<notin> A"
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   702
  obtains y where "v = f x y" and "fold_graph f z A y"
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   703
using assms by (auto dest: fold_graph_insertE_aux [OF _ insertI1])
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   704
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   705
lemma fold_graph_determ:
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   706
  "fold_graph f z A x \<Longrightarrow> fold_graph f z A y \<Longrightarrow> y = x"
36045
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   707
proof (induct arbitrary: y set: fold_graph)
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   708
  case (insertI x A y v)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   709
  from \<open>fold_graph f z (insert x A) v\<close> and \<open>x \<notin> A\<close>
36045
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   710
  obtain y' where "v = f x y'" and "fold_graph f z A y'"
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   711
    by (rule fold_graph_insertE)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   712
  from \<open>fold_graph f z A y'\<close> have "y' = y" by (rule insertI)
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   713
  with \<open>v = f x y'\<close> show "v = f x y" by simp
36045
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   714
qed fast
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   715
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   716
lemma fold_equality:
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   717
  "fold_graph f z A y \<Longrightarrow> fold f z A = y"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   718
  by (cases "finite A") (auto simp add: fold_def intro: fold_graph_determ dest: fold_graph_finite)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   719
42272
a46a13b4be5f dropped unused lemmas; proper Isar proof
haftmann
parents: 42207
diff changeset
   720
lemma fold_graph_fold:
a46a13b4be5f dropped unused lemmas; proper Isar proof
haftmann
parents: 42207
diff changeset
   721
  assumes "finite A"
a46a13b4be5f dropped unused lemmas; proper Isar proof
haftmann
parents: 42207
diff changeset
   722
  shows "fold_graph f z A (fold f z A)"
a46a13b4be5f dropped unused lemmas; proper Isar proof
haftmann
parents: 42207
diff changeset
   723
proof -
a46a13b4be5f dropped unused lemmas; proper Isar proof
haftmann
parents: 42207
diff changeset
   724
  from assms have "\<exists>x. fold_graph f z A x" by (rule finite_imp_fold_graph)
a46a13b4be5f dropped unused lemmas; proper Isar proof
haftmann
parents: 42207
diff changeset
   725
  moreover note fold_graph_determ
a46a13b4be5f dropped unused lemmas; proper Isar proof
haftmann
parents: 42207
diff changeset
   726
  ultimately have "\<exists>!x. fold_graph f z A x" by (rule ex_ex1I)
a46a13b4be5f dropped unused lemmas; proper Isar proof
haftmann
parents: 42207
diff changeset
   727
  then have "fold_graph f z A (The (fold_graph f z A))" by (rule theI')
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   728
  with assms show ?thesis by (simp add: fold_def)
42272
a46a13b4be5f dropped unused lemmas; proper Isar proof
haftmann
parents: 42207
diff changeset
   729
qed
36045
b846881928ea simplify fold_graph proofs
huffman
parents: 35831
diff changeset
   730
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61778
diff changeset
   731
text \<open>The base case for \<open>fold\<close>:\<close>
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   732
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   733
lemma (in -) fold_infinite [simp]:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   734
  assumes "\<not> finite A"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   735
  shows "fold f z A = z"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   736
  using assms by (auto simp add: fold_def)
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   737
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   738
lemma (in -) fold_empty [simp]:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   739
  "fold f z {} = z"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   740
  by (auto simp add: fold_def)
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   741
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   742
text\<open>The various recursion equations for @{const fold}:\<close>
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   743
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   744
lemma fold_insert [simp]:
42875
d1aad0957eb2 tuned proofs
haftmann
parents: 42873
diff changeset
   745
  assumes "finite A" and "x \<notin> A"
d1aad0957eb2 tuned proofs
haftmann
parents: 42873
diff changeset
   746
  shows "fold f z (insert x A) = f x (fold f z A)"
d1aad0957eb2 tuned proofs
haftmann
parents: 42873
diff changeset
   747
proof (rule fold_equality)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   748
  fix z
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   749
  from \<open>finite A\<close> have "fold_graph f z A (fold f z A)" by (rule fold_graph_fold)
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   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)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   751
  then show "fold_graph f z (insert x A) (f x (fold f z A))" by simp
42875
d1aad0957eb2 tuned proofs
haftmann
parents: 42873
diff changeset
   752
qed
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   753
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   754
declare (in -) empty_fold_graphE [rule del] fold_graph.intros [rule del]
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61778
diff changeset
   755
  \<comment> \<open>No more proofs involve these.\<close>
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   756
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   757
lemma fold_fun_left_comm:
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   758
  "finite A \<Longrightarrow> f x (fold f z A) = fold f (f x z) A"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   759
proof (induct rule: finite_induct)
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   760
  case empty then show ?case by simp
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   761
next
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   762
  case (insert y A) then show ?case
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   763
    by (simp add: fun_left_comm [of x])
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   764
qed
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   765
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   766
lemma fold_insert2:
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   767
  "finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> fold f z (insert x A)  = fold f (f x z) A"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   768
  by (simp add: fold_fun_left_comm)
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   769
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   770
lemma fold_rec:
42875
d1aad0957eb2 tuned proofs
haftmann
parents: 42873
diff changeset
   771
  assumes "finite A" and "x \<in> A"
d1aad0957eb2 tuned proofs
haftmann
parents: 42873
diff changeset
   772
  shows "fold f z A = f x (fold f z (A - {x}))"
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   773
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   774
  have A: "A = insert x (A - {x})" using \<open>x \<in> A\<close> by blast
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   775
  then have "fold f z A = fold f z (insert x (A - {x}))" by simp
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   776
  also have "\<dots> = f x (fold f z (A - {x}))"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   777
    by (rule fold_insert) (simp add: \<open>finite A\<close>)+
15535
nipkow
parents: 15532
diff changeset
   778
  finally show ?thesis .
nipkow
parents: 15532
diff changeset
   779
qed
nipkow
parents: 15532
diff changeset
   780
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   781
lemma fold_insert_remove:
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   782
  assumes "finite A"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   783
  shows "fold f z (insert x A) = f x (fold f z (A - {x}))"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   784
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   785
  from \<open>finite A\<close> have "finite (insert x A)" by auto
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   786
  moreover have "x \<in> insert x A" by auto
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   787
  ultimately have "fold f z (insert x A) = f x (fold f z (insert x A - {x}))"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   788
    by (rule fold_rec)
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   789
  then show ?thesis by simp
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   790
qed
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   791
57598
56ed992b6d65 add lemma
Andreas Lochbihler
parents: 57447
diff changeset
   792
lemma fold_set_union_disj:
56ed992b6d65 add lemma
Andreas Lochbihler
parents: 57447
diff changeset
   793
  assumes "finite A" "finite B" "A \<inter> B = {}"
56ed992b6d65 add lemma
Andreas Lochbihler
parents: 57447
diff changeset
   794
  shows "Finite_Set.fold f z (A \<union> B) = Finite_Set.fold f (Finite_Set.fold f z A) B"
56ed992b6d65 add lemma
Andreas Lochbihler
parents: 57447
diff changeset
   795
using assms(2,1,3) by induction simp_all
56ed992b6d65 add lemma
Andreas Lochbihler
parents: 57447
diff changeset
   796
51598
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   797
end
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   798
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   799
text\<open>Other properties of @{const fold}:\<close>
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
   800
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
   801
lemma fold_image:
51598
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   802
  assumes "inj_on g A"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   803
  shows "fold f z (g ` A) = fold (f \<circ> g) z A"
51598
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   804
proof (cases "finite A")
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   805
  case False with assms show ?thesis by (auto dest: finite_imageD simp add: fold_def)
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   806
next
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   807
  case True
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   808
  have "fold_graph f z (g ` A) = fold_graph (f \<circ> g) z A"
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   809
  proof
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   810
    fix w
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   811
    show "fold_graph f z (g ` A) w \<longleftrightarrow> fold_graph (f \<circ> g) z A w" (is "?P \<longleftrightarrow> ?Q")
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   812
    proof
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   813
      assume ?P then show ?Q using assms
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   814
      proof (induct "g ` A" w arbitrary: A)
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   815
        case emptyI then show ?case by (auto intro: fold_graph.emptyI)
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   816
      next
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   817
        case (insertI x A r B)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   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
51598
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   819
          "x' \<notin> A'" and [simp]: "B = insert x' A'" "x = g x'" "A = g ` A'"
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   820
          by (rule inj_img_insertE)
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   821
        from insertI.prems have "fold_graph (f o g) z A' r"
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   822
          by (auto intro: insertI.hyps)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   823
        with \<open>x' \<notin> A'\<close> have "fold_graph (f \<circ> g) z (insert x' A') ((f \<circ> g) x' r)"
51598
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   824
          by (rule fold_graph.insertI)
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   825
        then show ?case by simp
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   826
      qed
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   827
    next
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   828
      assume ?Q then show ?P using assms
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   829
      proof induct
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   830
        case emptyI thus ?case by (auto intro: fold_graph.emptyI)
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   831
      next
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   832
        case (insertI x A r)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   833
        from \<open>x \<notin> A\<close> insertI.prems have "g x \<notin> g ` A" by auto
51598
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   834
        moreover from insertI have "fold_graph f z (g ` A) r" by simp
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   835
        ultimately have "fold_graph f z (insert (g x) (g ` A)) (f (g x) r)"
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   836
          by (rule fold_graph.insertI)
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   837
        then show ?case by simp
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   838
      qed
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   839
    qed
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   840
  qed
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   841
  with True assms show ?thesis by (auto simp add: fold_def)
5dbe537087aa generalized lemma fold_image thanks to Peter Lammich
haftmann
parents: 51546
diff changeset
   842
qed
15392
290bc97038c7 First step in reorganizing Finite_Set
nipkow
parents: 15376
diff changeset
   843
49724
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   844
lemma fold_cong:
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   845
  assumes "comp_fun_commute f" "comp_fun_commute g"
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   846
  assumes "finite A" and cong: "\<And>x. x \<in> A \<Longrightarrow> f x = g x"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   847
    and "s = t" and "A = B"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   848
  shows "fold f s A = fold g t B"
49724
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   849
proof -
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   850
  have "fold f s A = fold g s A"  
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   851
  using \<open>finite A\<close> cong proof (induct A)
49724
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   852
    case empty then show ?case by simp
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   853
  next
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   854
    case (insert x A)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   855
    interpret f: comp_fun_commute f by (fact \<open>comp_fun_commute f\<close>)
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   856
    interpret g: comp_fun_commute g by (fact \<open>comp_fun_commute g\<close>)
49724
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   857
    from insert show ?case by simp
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   858
  qed
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   859
  with assms show ?thesis by simp
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   860
qed
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   861
a5842f026d4c congruence rule for Finite_Set.fold
haftmann
parents: 49723
diff changeset
   862
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   863
text \<open>A simplified version for idempotent functions:\<close>
15480
cb3612cc41a3 renamed a few vars, added a lemma
nipkow
parents: 15479
diff changeset
   864
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   865
locale comp_fun_idem = comp_fun_commute +
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   866
  assumes comp_fun_idem: "f x \<circ> f x = f x"
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   867
begin
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   868
42869
43b0f61f56d0 use point-free characterization for locale fun_left_comm_idem
haftmann
parents: 42809
diff changeset
   869
lemma fun_left_idem: "f x (f x z) = f x z"
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   870
  using comp_fun_idem by (simp add: fun_eq_iff)
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   871
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   872
lemma fold_insert_idem:
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   873
  assumes fin: "finite A"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   874
  shows "fold f z (insert x A)  = f x (fold f z A)"
15480
cb3612cc41a3 renamed a few vars, added a lemma
nipkow
parents: 15479
diff changeset
   875
proof cases
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   876
  assume "x \<in> A"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   877
  then obtain B where "A = insert x B" and "x \<notin> B" by (rule set_insert)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   878
  then show ?thesis using assms by (simp add: comp_fun_idem fun_left_idem)
15480
cb3612cc41a3 renamed a few vars, added a lemma
nipkow
parents: 15479
diff changeset
   879
next
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   880
  assume "x \<notin> A" then show ?thesis using assms by simp
15480
cb3612cc41a3 renamed a few vars, added a lemma
nipkow
parents: 15479
diff changeset
   881
qed
cb3612cc41a3 renamed a few vars, added a lemma
nipkow
parents: 15479
diff changeset
   882
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   883
declare fold_insert [simp del] fold_insert_idem [simp]
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   884
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   885
lemma fold_insert_idem2:
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28823
diff changeset
   886
  "finite A \<Longrightarrow> fold f z (insert x A) = fold f (f x z) A"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   887
  by (simp add: fold_fun_left_comm)
15484
2636ec211ec8 fold and fol1 changes
nipkow
parents: 15483
diff changeset
   888
26041
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   889
end
c2e15e65165f locales ACf, ACIf, ACIfSL and ACIfSLlin have been abandoned in favour of the existing algebraic classes ab_semigroup_mult, ab_semigroup_idem_mult, lower_semilattice (resp. uper_semilattice) and linorder
haftmann
parents: 25571
diff changeset
   890
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   891
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61778
diff changeset
   892
subsubsection \<open>Liftings to \<open>comp_fun_commute\<close> etc.\<close>
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   893
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   894
lemma (in comp_fun_commute) comp_comp_fun_commute:
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   895
  "comp_fun_commute (f \<circ> g)"
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   896
proof
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   897
qed (simp_all add: comp_fun_commute)
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   898
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   899
lemma (in comp_fun_idem) comp_comp_fun_idem:
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   900
  "comp_fun_idem (f \<circ> g)"
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   901
  by (rule comp_fun_idem.intro, rule comp_comp_fun_commute, unfold_locales)
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   902
    (simp_all add: comp_fun_idem)
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   903
49723
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   904
lemma (in comp_fun_commute) comp_fun_commute_funpow:
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   905
  "comp_fun_commute (\<lambda>x. f x ^^ g x)"
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   906
proof
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   907
  fix y x
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   908
  show "f y ^^ g y \<circ> f x ^^ g x = f x ^^ g x \<circ> f y ^^ g y"
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   909
  proof (cases "x = y")
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   910
    case True then show ?thesis by simp
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   911
  next
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   912
    case False show ?thesis
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   913
    proof (induct "g x" arbitrary: g)
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   914
      case 0 then show ?case by simp
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   915
    next
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   916
      case (Suc n g)
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   917
      have hyp1: "f y ^^ g y \<circ> f x = f x \<circ> f y ^^ g y"
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   918
      proof (induct "g y" arbitrary: g)
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   919
        case 0 then show ?case by simp
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   920
      next
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   921
        case (Suc n g)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62618
diff changeset
   922
        define h where "h z = g z - 1" for z
49723
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   923
        with Suc have "n = h y" by simp
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   924
        with Suc have hyp: "f y ^^ h y \<circ> f x = f x \<circ> f y ^^ h y"
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   925
          by auto
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   926
        from Suc h_def have "g y = Suc (h y)" by simp
49739
13aa6d8268ec consolidated names of theorems on composition;
haftmann
parents: 49724
diff changeset
   927
        then show ?case by (simp add: comp_assoc hyp)
49723
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   928
          (simp add: o_assoc comp_fun_commute)
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   929
      qed
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62618
diff changeset
   930
      define h where "h z = (if z = x then g x - 1 else g z)" for z
49723
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   931
      with Suc have "n = h x" by simp
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   932
      with Suc have "f y ^^ h y \<circ> f x ^^ h x = f x ^^ h x \<circ> f y ^^ h y"
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   933
        by auto
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   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
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   935
      from Suc h_def have "g x = Suc (h x)" by simp
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   936
      then show ?case by (simp del: funpow.simps add: funpow_Suc_right o_assoc hyp2)
49739
13aa6d8268ec consolidated names of theorems on composition;
haftmann
parents: 49724
diff changeset
   937
        (simp add: comp_assoc hyp1)
49723
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   938
    qed
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   939
  qed
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   940
qed
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   941
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   942
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   943
subsubsection \<open>Expressing set operations via @{const fold}\<close>
49723
bbc2942ba09f alternative simplification of ^^ to the righthand side;
haftmann
parents: 48891
diff changeset
   944
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   945
lemma comp_fun_commute_const:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   946
  "comp_fun_commute (\<lambda>_. f)"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   947
proof
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   948
qed rule
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   949
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   950
lemma comp_fun_idem_insert:
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   951
  "comp_fun_idem insert"
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   952
proof
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   953
qed auto
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   954
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   955
lemma comp_fun_idem_remove:
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
   956
  "comp_fun_idem Set.remove"
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   957
proof
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   958
qed auto
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
   959
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   960
lemma (in semilattice_inf) comp_fun_idem_inf:
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   961
  "comp_fun_idem inf"
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   962
proof
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   963
qed (auto simp add: inf_left_commute)
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   964
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   965
lemma (in semilattice_sup) comp_fun_idem_sup:
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   966
  "comp_fun_idem sup"
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   967
proof
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   968
qed (auto simp add: sup_left_commute)
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
   969
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   970
lemma union_fold_insert:
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   971
  assumes "finite A"
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   972
  shows "A \<union> B = fold insert B A"
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   973
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
   974
  interpret comp_fun_idem insert by (fact comp_fun_idem_insert)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   975
  from \<open>finite A\<close> show ?thesis by (induct A arbitrary: B) simp_all
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   976
qed
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
   977
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   978
lemma minus_fold_remove:
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   979
  assumes "finite A"
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
   980
  shows "B - A = fold Set.remove B A"
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   981
proof -
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
   982
  interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
   983
  from \<open>finite A\<close> have "fold Set.remove B A = B - A" by (induct A arbitrary: B) auto
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
   984
  then show ?thesis ..
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   985
qed
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
   986
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   987
lemma comp_fun_commute_filter_fold:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
   988
  "comp_fun_commute (\<lambda>x A'. if P x then Set.insert x A' else A')"
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
   989
proof - 
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
   990
  interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert)
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
   991
  show ?thesis by standard (auto simp: fun_eq_iff)
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
   992
qed
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
   993
49758
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
   994
lemma Set_filter_fold:
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
   995
  assumes "finite A"
49758
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
   996
  shows "Set.filter P A = fold (\<lambda>x A'. if P x then Set.insert x A' else A') {} A"
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
   997
using assms
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
   998
by (induct A) 
49758
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
   999
  (auto simp add: Set.filter_def comp_fun_commute.fold_insert[OF comp_fun_commute_filter_fold])
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
  1000
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
  1001
lemma inter_Set_filter:     
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
  1002
  assumes "finite B"
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
  1003
  shows "A \<inter> B = Set.filter (\<lambda>x. x \<in> A) B"
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
  1004
using assms 
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
  1005
by (induct B) (auto simp: Set.filter_def)
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1006
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1007
lemma image_fold_insert:
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1008
  assumes "finite A"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1009
  shows "image f A = fold (\<lambda>k A. Set.insert (f k) A) {} A"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1010
using assms
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1011
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1012
  interpret comp_fun_commute "\<lambda>k A. Set.insert (f k) A" by standard auto
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1013
  show ?thesis using assms by (induct A) auto
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1014
qed
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1015
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1016
lemma Ball_fold:
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1017
  assumes "finite A"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1018
  shows "Ball A P = fold (\<lambda>k s. s \<and> P k) True A"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1019
using assms
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1020
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1021
  interpret comp_fun_commute "\<lambda>k s. s \<and> P k" by standard auto
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1022
  show ?thesis using assms by (induct A) auto
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1023
qed
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1024
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1025
lemma Bex_fold:
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1026
  assumes "finite A"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1027
  shows "Bex A P = fold (\<lambda>k s. s \<or> P k) False A"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1028
using assms
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1029
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1030
  interpret comp_fun_commute "\<lambda>k s. s \<or> P k" by standard auto
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1031
  show ?thesis using assms by (induct A) auto
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1032
qed
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1033
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1034
lemma comp_fun_commute_Pow_fold: 
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1035
  "comp_fun_commute (\<lambda>x A. A \<union> Set.insert x ` A)" 
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1036
  by (clarsimp simp: fun_eq_iff comp_fun_commute_def) blast
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1037
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1038
lemma Pow_fold:
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1039
  assumes "finite A"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1040
  shows "Pow A = fold (\<lambda>x A. A \<union> Set.insert x ` A) {{}} A"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1041
using assms
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1042
proof -
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1043
  interpret comp_fun_commute "\<lambda>x A. A \<union> Set.insert x ` A" by (rule comp_fun_commute_Pow_fold)
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1044
  show ?thesis using assms by (induct A) (auto simp: Pow_insert)
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1045
qed
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1046
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1047
lemma fold_union_pair:
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1048
  assumes "finite B"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1049
  shows "(\<Union>y\<in>B. {(x, y)}) \<union> A = fold (\<lambda>y. Set.insert (x, y)) A B"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1050
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1051
  interpret comp_fun_commute "\<lambda>y. Set.insert (x, y)" by standard auto
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1052
  show ?thesis using assms  by (induct B arbitrary: A) simp_all
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1053
qed
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1054
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1055
lemma comp_fun_commute_product_fold: 
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1056
  assumes "finite B"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1057
  shows "comp_fun_commute (\<lambda>x z. fold (\<lambda>y. Set.insert (x, y)) z B)" 
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1058
  by standard (auto simp: fold_union_pair[symmetric] assms)
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1059
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1060
lemma product_fold:
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1061
  assumes "finite A"
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1062
  assumes "finite B"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1063
  shows "A \<times> B = fold (\<lambda>x z. fold (\<lambda>y. Set.insert (x, y)) z B) {} A"
48619
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1064
using assms unfolding Sigma_def 
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1065
by (induct A) 
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1066
  (simp_all add: comp_fun_commute.fold_insert[OF comp_fun_commute_product_fold] fold_union_pair)
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1067
558e4e77ce69 more set operations expressed by Finite_Set.fold
kuncar
parents: 48175
diff changeset
  1068
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1069
context complete_lattice
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1070
begin
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1071
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1072
lemma inf_Inf_fold_inf:
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1073
  assumes "finite A"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1074
  shows "inf (Inf A) B = fold inf B A"
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1075
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
  1076
  interpret comp_fun_idem inf by (fact comp_fun_idem_inf)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1077
  from \<open>finite A\<close> fold_fun_left_comm show ?thesis by (induct A arbitrary: B)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1078
    (simp_all add: inf_commute fun_eq_iff)
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1079
qed
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1080
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1081
lemma sup_Sup_fold_sup:
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1082
  assumes "finite A"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1083
  shows "sup (Sup A) B = fold sup B A"
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1084
proof -
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
  1085
  interpret comp_fun_idem sup by (fact comp_fun_idem_sup)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1086
  from \<open>finite A\<close> fold_fun_left_comm show ?thesis by (induct A arbitrary: B)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1087
    (simp_all add: sup_commute fun_eq_iff)
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1088
qed
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1089
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1090
lemma Inf_fold_inf:
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1091
  assumes "finite A"
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1092
  shows "Inf A = fold inf top A"
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1093
  using assms inf_Inf_fold_inf [of A top] by (simp add: inf_absorb2)
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1094
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1095
lemma Sup_fold_sup:
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1096
  assumes "finite A"
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1097
  shows "Sup A = fold sup bot A"
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1098
  using assms sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2)
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1099
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1100
lemma inf_INF_fold_inf:
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1101
  assumes "finite A"
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56166
diff changeset
  1102
  shows "inf B (INFIMUM A f) = fold (inf \<circ> f) B A" (is "?inf = ?fold") 
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1103
proof (rule sym)
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
  1104
  interpret comp_fun_idem inf by (fact comp_fun_idem_inf)
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
  1105
  interpret comp_fun_idem "inf \<circ> f" by (fact comp_comp_fun_idem)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1106
  from \<open>finite A\<close> show "?fold = ?inf"
42869
43b0f61f56d0 use point-free characterization for locale fun_left_comm_idem
haftmann
parents: 42809
diff changeset
  1107
    by (induct A arbitrary: B)
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56154
diff changeset
  1108
      (simp_all add: inf_left_commute)
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1109
qed
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1110
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1111
lemma sup_SUP_fold_sup:
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1112
  assumes "finite A"
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56166
diff changeset
  1113
  shows "sup B (SUPREMUM A f) = fold (sup \<circ> f) B A" (is "?sup = ?fold") 
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1114
proof (rule sym)
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
  1115
  interpret comp_fun_idem sup by (fact comp_fun_idem_sup)
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
  1116
  interpret comp_fun_idem "sup \<circ> f" by (fact comp_comp_fun_idem)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1117
  from \<open>finite A\<close> show "?fold = ?sup"
42869
43b0f61f56d0 use point-free characterization for locale fun_left_comm_idem
haftmann
parents: 42809
diff changeset
  1118
    by (induct A arbitrary: B)
56166
9a241bc276cd normalising simp rules for compound operators
haftmann
parents: 56154
diff changeset
  1119
      (simp_all add: sup_left_commute)
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1120
qed
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1121
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1122
lemma INF_fold_inf:
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1123
  assumes "finite A"
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56166
diff changeset
  1124
  shows "INFIMUM A f = fold (inf \<circ> f) top A"
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1125
  using assms inf_INF_fold_inf [of A top] by simp
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1126
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1127
lemma SUP_fold_sup:
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1128
  assumes "finite A"
56218
1c3f1f2431f9 elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents: 56166
diff changeset
  1129
  shows "SUPREMUM A f = fold (sup \<circ> f) bot A"
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1130
  using assms sup_SUP_fold_sup [of A bot] by simp
31992
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1131
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1132
end
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1133
f8aed98faae7 More about gcd/lcm, and some cleaning up
nipkow
parents: 31916
diff changeset
  1134
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1135
subsection \<open>Locales as mini-packages for fold operations\<close>
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33960
diff changeset
  1136
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1137
subsubsection \<open>The natural case\<close>
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1138
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1139
locale folding =
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1140
  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1141
  fixes z :: "'b"
42871
1c0b99f950d9 names of fold_set locales resemble name of characteristic property more closely
haftmann
parents: 42869
diff changeset
  1142
  assumes comp_fun_commute: "f y \<circ> f x = f x \<circ> f y"
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1143
begin
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1144
54870
1b5f2485757b prefix disambiguation
haftmann
parents: 54867
diff changeset
  1145
interpretation fold?: comp_fun_commute f
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1146
  by standard (insert comp_fun_commute, simp add: fun_eq_iff)
54867
c21a2465cac1 prefer ephemeral interpretation over interpretation in proof contexts;
haftmann
parents: 54611
diff changeset
  1147
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1148
definition F :: "'a set \<Rightarrow> 'b"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1149
where
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1150
  eq_fold: "F A = fold f z A"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1151
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1152
lemma empty [simp]:"F {} = z"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1153
  by (simp add: eq_fold)
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1154
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1155
lemma infinite [simp]: "\<not> finite A \<Longrightarrow> F A = z"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1156
  by (simp add: eq_fold)
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1157
 
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1158
lemma insert [simp]:
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1159
  assumes "finite A" and "x \<notin> A"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1160
  shows "F (insert x A) = f x (F A)"
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1161
proof -
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1162
  from fold_insert assms
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1163
  have "fold f z (insert x A) = f x (fold f z A)" by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1164
  with \<open>finite A\<close> show ?thesis by (simp add: eq_fold fun_eq_iff)
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1165
qed
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1166
 
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1167
lemma remove:
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1168
  assumes "finite A" and "x \<in> A"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1169
  shows "F A = f x (F (A - {x}))"
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1170
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1171
  from \<open>x \<in> A\<close> obtain B where A: "A = insert x B" and "x \<notin> B"
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1172
    by (auto dest: mk_disjoint_insert)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1173
  moreover from \<open>finite A\<close> A have "finite B" by simp
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1174
  ultimately show ?thesis by simp
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1175
qed
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1176
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1177
lemma insert_remove:
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1178
  assumes "finite A"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1179
  shows "F (insert x A) = f x (F (A - {x}))"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1180
  using assms by (cases "x \<in> A") (simp_all add: remove insert_absorb)
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1181
34007
aea892559fc5 tuned lattices theory fragements; generlized some lemmas from sets to lattices
haftmann
parents: 33960
diff changeset
  1182
end
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1183
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1184
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1185
subsubsection \<open>With idempotency\<close>
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1186
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1187
locale folding_idem = folding +
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1188
  assumes comp_fun_idem: "f x \<circ> f x = f x"
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1189
begin
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1190
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1191
declare insert [simp del]
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1192
54870
1b5f2485757b prefix disambiguation
haftmann
parents: 54867
diff changeset
  1193
interpretation fold?: comp_fun_idem f
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1194
  by standard (insert comp_fun_commute comp_fun_idem, simp add: fun_eq_iff)
54867
c21a2465cac1 prefer ephemeral interpretation over interpretation in proof contexts;
haftmann
parents: 54611
diff changeset
  1195
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1196
lemma insert_idem [simp]:
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1197
  assumes "finite A"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1198
  shows "F (insert x A) = f x (F A)"
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1199
proof -
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1200
  from fold_insert_idem assms
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1201
  have "fold f z (insert x A) = f x (fold f z A)" by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1202
  with \<open>finite A\<close> show ?thesis by (simp add: eq_fold fun_eq_iff)
35719
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1203
qed
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1204
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1205
end
99b6152aedf5 split off theory Big_Operators from theory Finite_Set
haftmann
parents: 35577
diff changeset
  1206
35817
d8b8527102f5 added locales folding_one_(idem); various streamlining and tuning
haftmann
parents: 35796
diff changeset
  1207
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1208
subsection \<open>Finite cardinality\<close>
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1209
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1210
text \<open>
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1211
  The traditional definition
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1212
  @{prop "card A \<equiv> LEAST n. EX f. A = {f i | i. i < n}"}
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1213
  is ugly to work with.
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1214
  But now that we have @{const fold} things are easy:
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1215
\<close>
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1216
61890
f6ded81f5690 abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents: 61810
diff changeset
  1217
global_interpretation card: folding "\<lambda>_. Suc" 0
61778
9f4f0dc7be2c modernized
haftmann
parents: 61762
diff changeset
  1218
  defines card = "folding.F (\<lambda>_. Suc) 0"
9f4f0dc7be2c modernized
haftmann
parents: 61762
diff changeset
  1219
  by standard rule
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1220
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1221
lemma card_infinite:
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1222
  "\<not> finite A \<Longrightarrow> card A = 0"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1223
  by (fact card.infinite)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1224
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1225
lemma card_empty:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1226
  "card {} = 0"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1227
  by (fact card.empty)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1228
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1229
lemma card_insert_disjoint:
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1230
  "finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> card (insert x A) = Suc (card A)"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1231
  by (fact card.insert)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1232
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1233
lemma card_insert_if:
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1234
  "finite A \<Longrightarrow> card (insert x A) = (if x \<in> A then card A else Suc (card A))"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1235
  by auto (simp add: card.insert_remove card.remove)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1236
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1237
lemma card_ge_0_finite:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1238
  "card A > 0 \<Longrightarrow> finite A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1239
  by (rule ccontr) simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1240
54148
c8cc5ab4a863 killed more "no_atp"s
blanchet
parents: 54147
diff changeset
  1241
lemma card_0_eq [simp]:
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1242
  "finite A \<Longrightarrow> card A = 0 \<longleftrightarrow> A = {}"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1243
  by (auto dest: mk_disjoint_insert)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1244
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1245
lemma finite_UNIV_card_ge_0:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1246
  "finite (UNIV :: 'a set) \<Longrightarrow> card (UNIV :: 'a set) > 0"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1247
  by (rule ccontr) simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1248
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1249
lemma card_eq_0_iff:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1250
  "card A = 0 \<longleftrightarrow> A = {} \<or> \<not> finite A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1251
  by auto
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1252
63365
5340fb6633d0 more theorems
haftmann
parents: 63099
diff changeset
  1253
lemma card_range_greater_zero:
5340fb6633d0 more theorems
haftmann
parents: 63099
diff changeset
  1254
  "finite (range f) \<Longrightarrow> card (range f) > 0"
5340fb6633d0 more theorems
haftmann
parents: 63099
diff changeset
  1255
  by (rule ccontr) (simp add: card_eq_0_iff)
5340fb6633d0 more theorems
haftmann
parents: 63099
diff changeset
  1256
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1257
lemma card_gt_0_iff:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1258
  "0 < card A \<longleftrightarrow> A \<noteq> {} \<and> finite A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1259
  by (simp add: neq0_conv [symmetric] card_eq_0_iff) 
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1260
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1261
lemma card_Suc_Diff1:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1262
  "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> Suc (card (A - {x})) = card A"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1263
apply(rule_tac t = A in insert_Diff [THEN subst], assumption)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1264
apply(simp del:insert_Diff_single)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1265
done
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1266
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60758
diff changeset
  1267
lemma card_insert_le_m1: "n>0 \<Longrightarrow> card y \<le> n-1 \<Longrightarrow> card (insert x y) \<le> n"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60758
diff changeset
  1268
  apply (cases "finite y")
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60758
diff changeset
  1269
  apply (cases "x \<in> y")
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60758
diff changeset
  1270
  apply (auto simp: insert_absorb)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60758
diff changeset
  1271
  done
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60758
diff changeset
  1272
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1273
lemma card_Diff_singleton:
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1274
  "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> card (A - {x}) = card A - 1"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1275
  by (simp add: card_Suc_Diff1 [symmetric])
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1276
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1277
lemma card_Diff_singleton_if:
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1278
  "finite A \<Longrightarrow> card (A - {x}) = (if x \<in> A then card A - 1 else card A)"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1279
  by (simp add: card_Diff_singleton)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1280
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1281
lemma card_Diff_insert[simp]:
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1282
  assumes "finite A" and "a \<in> A" and "a \<notin> B"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1283
  shows "card (A - insert a B) = card (A - B) - 1"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1284
proof -
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1285
  have "A - insert a B = (A - B) - {a}" using assms by blast
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1286
  then show ?thesis using assms by(simp add: card_Diff_singleton)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1287
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1288
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1289
lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1290
  by (fact card.insert_remove)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1291
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1292
lemma card_insert_le: "finite A ==> card A <= card (insert x A)"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1293
by (simp add: card_insert_if)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1294
41987
4ad8f1dc2e0b added lemmas
nipkow
parents: 41657
diff changeset
  1295
lemma card_Collect_less_nat[simp]: "card{i::nat. i < n} = n"
4ad8f1dc2e0b added lemmas
nipkow
parents: 41657
diff changeset
  1296
by (induct n) (simp_all add:less_Suc_eq Collect_disj_eq)
4ad8f1dc2e0b added lemmas
nipkow
parents: 41657
diff changeset
  1297
41988
c2583bbb92f5 tuned lemma
nipkow
parents: 41987
diff changeset
  1298
lemma card_Collect_le_nat[simp]: "card{i::nat. i <= n} = Suc n"
41987
4ad8f1dc2e0b added lemmas
nipkow
parents: 41657
diff changeset
  1299
using card_Collect_less_nat[of "Suc n"] by(simp add: less_Suc_eq_le)
4ad8f1dc2e0b added lemmas
nipkow
parents: 41657
diff changeset
  1300
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1301
lemma card_mono:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1302
  assumes "finite B" and "A \<subseteq> B"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1303
  shows "card A \<le> card B"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1304
proof -
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1305
  from assms have "finite A" by (auto intro: finite_subset)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1306
  then show ?thesis using assms proof (induct A arbitrary: B)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1307
    case empty then show ?case by simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1308
  next
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1309
    case (insert x A)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1310
    then have "x \<in> B" by simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1311
    from insert have "A \<subseteq> B - {x}" and "finite (B - {x})" by auto
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1312
    with insert.hyps have "card A \<le> card (B - {x})" by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  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)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1314
  qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1315
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1316
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1317
lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)"
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
  1318
apply (induct rule: finite_induct)
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
  1319
apply simp
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
  1320
apply clarify
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1321
apply (subgoal_tac "finite A & A - {x} <= F")
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1322
 prefer 2 apply (blast intro: finite_subset, atomize)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1323
apply (drule_tac x = "A - {x}" in spec)
62390
842917225d56 more canonical names
nipkow
parents: 62093
diff changeset
  1324
apply (simp add: card_Diff_singleton_if split add: if_split_asm)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1325
apply (case_tac "card A", auto)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1326
done
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1327
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1328
lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1329
apply (simp add: psubset_eq linorder_not_le [symmetric])
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1330
apply (blast dest: card_seteq)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1331
done
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1332
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1333
lemma card_Un_Int:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1334
  assumes "finite A" and "finite B"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1335
  shows "card A + card B = card (A \<union> B) + card (A \<inter> B)"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1336
using assms proof (induct A)
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1337
  case empty then show ?case by simp
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1338
next
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1339
 case (insert x A) then show ?case
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1340
    by (auto simp add: insert_absorb Int_insert_left)
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1341
qed
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1342
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1343
lemma card_Un_disjoint:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1344
  assumes "finite A" and "finite B"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1345
  assumes "A \<inter> B = {}"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1346
  shows "card (A \<union> B) = card A + card B"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1347
using assms card_Un_Int [of A B] by simp
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1348
59336
a95b6f608a73 added lemma
nipkow
parents: 58889
diff changeset
  1349
lemma card_Un_le: "card (A \<union> B) \<le> card A + card B"
a95b6f608a73 added lemma
nipkow
parents: 58889
diff changeset
  1350
apply(cases "finite A")
a95b6f608a73 added lemma
nipkow
parents: 58889
diff changeset
  1351
 apply(cases "finite B")
a95b6f608a73 added lemma
nipkow
parents: 58889
diff changeset
  1352
  using le_iff_add card_Un_Int apply blast
a95b6f608a73 added lemma
nipkow
parents: 58889
diff changeset
  1353
 apply simp
a95b6f608a73 added lemma
nipkow
parents: 58889
diff changeset
  1354
apply simp
a95b6f608a73 added lemma
nipkow
parents: 58889
diff changeset
  1355
done
a95b6f608a73 added lemma
nipkow
parents: 58889
diff changeset
  1356
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1357
lemma card_Diff_subset:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1358
  assumes "finite B" and "B \<subseteq> A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1359
  shows "card (A - B) = card A - card B"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1360
proof (cases "finite A")
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1361
  case False with assms show ?thesis by simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1362
next
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1363
  case True with assms show ?thesis by (induct B arbitrary: A) simp_all
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1364
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1365
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1366
lemma card_Diff_subset_Int:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1367
  assumes AB: "finite (A \<inter> B)" shows "card (A - B) = card A - card (A \<inter> B)"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1368
proof -
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1369
  have "A - B = A - A \<inter> B" by auto
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1370
  thus ?thesis
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1371
    by (simp add: card_Diff_subset AB) 
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1372
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1373
40716
a92d744bca5f new lemma
nipkow
parents: 40703
diff changeset
  1374
lemma diff_card_le_card_Diff:
a92d744bca5f new lemma
nipkow
parents: 40703
diff changeset
  1375
assumes "finite B" shows "card A - card B \<le> card(A - B)"
a92d744bca5f new lemma
nipkow
parents: 40703
diff changeset
  1376
proof-
a92d744bca5f new lemma
nipkow
parents: 40703
diff changeset
  1377
  have "card A - card B \<le> card A - card (A \<inter> B)"
a92d744bca5f new lemma
nipkow
parents: 40703
diff changeset
  1378
    using card_mono[OF assms Int_lower2, of A] by arith
a92d744bca5f new lemma
nipkow
parents: 40703
diff changeset
  1379
  also have "\<dots> = card(A-B)" using assms by(simp add: card_Diff_subset_Int)
a92d744bca5f new lemma
nipkow
parents: 40703
diff changeset
  1380
  finally show ?thesis .
a92d744bca5f new lemma
nipkow
parents: 40703
diff changeset
  1381
qed
a92d744bca5f new lemma
nipkow
parents: 40703
diff changeset
  1382
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1383
lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1384
apply (rule Suc_less_SucD)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1385
apply (simp add: card_Suc_Diff1 del:card_Diff_insert)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1386
done
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1387
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1388
lemma card_Diff2_less:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1389
  "finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1390
apply (case_tac "x = y")
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1391
 apply (simp add: card_Diff1_less del:card_Diff_insert)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1392
apply (rule less_trans)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1393
 prefer 2 apply (auto intro!: card_Diff1_less simp del:card_Diff_insert)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1394
done
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1395
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1396
lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1397
apply (case_tac "x : A")
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1398
 apply (simp_all add: card_Diff1_less less_imp_le)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1399
done
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1400
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1401
lemma card_psubset: "finite B ==> A \<subseteq> B ==> card A < card B ==> A < B"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1402
by (erule psubsetI, blast)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1403
54413
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1404
lemma card_le_inj:
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1405
  assumes fA: "finite A"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1406
    and fB: "finite B"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1407
    and c: "card A \<le> card B"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1408
  shows "\<exists>f. f ` A \<subseteq> B \<and> inj_on f A"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1409
  using fA fB c
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1410
proof (induct arbitrary: B rule: finite_induct)
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1411
  case empty
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1412
  then show ?case by simp
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1413
next
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1414
  case (insert x s t)
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1415
  then show ?case
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1416
  proof (induct rule: finite_induct[OF "insert.prems"(1)])
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1417
    case 1
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1418
    then show ?case by simp
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1419
  next
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1420
    case (2 y t)
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1421
    from "2.prems"(1,2,5) "2.hyps"(1,2) have cst: "card s \<le> card t"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1422
      by simp
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1423
    from "2.prems"(3) [OF "2.hyps"(1) cst]
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1424
    obtain f where "f ` s \<subseteq> t" "inj_on f s"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1425
      by blast
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1426
    with "2.prems"(2) "2.hyps"(2) show ?case
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1427
      apply -
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1428
      apply (rule exI[where x = "\<lambda>z. if z = x then y else f z"])
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1429
      apply (auto simp add: inj_on_def)
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1430
      done
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1431
  qed
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1432
qed
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1433
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1434
lemma card_subset_eq:
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1435
  assumes fB: "finite B"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1436
    and AB: "A \<subseteq> B"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1437
    and c: "card A = card B"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1438
  shows "A = B"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1439
proof -
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1440
  from fB AB have fA: "finite A"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1441
    by (auto intro: finite_subset)
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1442
  from fA fB have fBA: "finite (B - A)"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1443
    by auto
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1444
  have e: "A \<inter> (B - A) = {}"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1445
    by blast
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1446
  have eq: "A \<union> (B - A) = B"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1447
    using AB by blast
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1448
  from card_Un_disjoint[OF fA fBA e, unfolded eq c] have "card (B - A) = 0"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1449
    by arith
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1450
  then have "B - A = {}"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1451
    unfolding card_eq_0_iff using fA fB by simp
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1452
  with AB show "A = B"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1453
    by blast
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1454
qed
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1455
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1456
lemma insert_partition:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  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>
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60303
diff changeset
  1458
  \<Longrightarrow> x \<inter> \<Union>F = {}"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1459
by auto
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1460
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1461
lemma finite_psubset_induct[consumes 1, case_names psubset]:
36079
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1462
  assumes fin: "finite A" 
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1463
  and     major: "\<And>A. finite A \<Longrightarrow> (\<And>B. B \<subset> A \<Longrightarrow> P B) \<Longrightarrow> P A" 
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1464
  shows "P A"
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1465
using fin
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1466
proof (induct A taking: card rule: measure_induct_rule)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1467
  case (less A)
36079
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1468
  have fin: "finite A" by fact
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1469
  have ih: "\<And>B. \<lbrakk>card B < card A; finite B\<rbrakk> \<Longrightarrow> P B" by fact
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1470
  { fix B 
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1471
    assume asm: "B \<subset> A"
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1472
    from asm have "card B < card A" using psubset_card_mono fin by blast
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1473
    moreover
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1474
    from asm have "B \<subseteq> A" by auto
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1475
    then have "finite B" using fin finite_subset by blast
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1476
    ultimately 
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1477
    have "P B" using ih by simp
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1478
  }
fa0e354e6a39 simplified induction case in finite_psubset_induct; tuned the proof that uses this induction principle
Christian Urban <urbanc@in.tum.de>
parents: 36045
diff changeset
  1479
  with fin show "P A" using major by blast
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1480
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1481
54413
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1482
lemma finite_induct_select[consumes 1, case_names empty select]:
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1483
  assumes "finite S"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1484
  assumes "P {}"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1485
  assumes select: "\<And>T. T \<subset> S \<Longrightarrow> P T \<Longrightarrow> \<exists>s\<in>S - T. P (insert s T)"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1486
  shows "P S"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1487
proof -
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1488
  have "0 \<le> card S" by simp
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1489
  then have "\<exists>T \<subseteq> S. card T = card S \<and> P T"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1490
  proof (induct rule: dec_induct)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1491
    case base with \<open>P {}\<close> show ?case
54413
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1492
      by (intro exI[of _ "{}"]) auto
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1493
  next
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1494
    case (step n)
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1495
    then obtain T where T: "T \<subseteq> S" "card T = n" "P T"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1496
      by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1497
    with \<open>n < card S\<close> have "T \<subset> S" "P T"
54413
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1498
      by auto
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1499
    with select[of T] obtain s where "s \<in> S" "s \<notin> T" "P (insert s T)"
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1500
      by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1501
    with step(2) T \<open>finite S\<close> show ?case
54413
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1502
      by (intro exI[of _ "insert s T"]) (auto dest: finite_subset)
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1503
  qed
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1504
  with \<open>finite S\<close> show "P S"
54413
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1505
    by (auto dest: card_subset_eq)
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1506
qed
88a036a95967 add finite_select_induct; move generic lemmas from MV_Analysis/Linear_Algebra to the HOL image
hoelzl
parents: 54148
diff changeset
  1507
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1508
lemma remove_induct [case_names empty infinite remove]:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1509
  assumes empty: "P ({} :: 'a set)" and infinite: "\<not>finite B \<Longrightarrow> P B"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  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"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1511
  shows "P B"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1512
proof (cases "finite B")
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1513
  assume "\<not>finite B"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1514
  thus ?thesis by (rule infinite)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1515
next
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1516
  define A where "A = B"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1517
  assume "finite B"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1518
  hence "finite A" "A \<subseteq> B" by (simp_all add: A_def)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1519
  thus "P A"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1520
  proof (induction "card A" arbitrary: A)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1521
    case 0
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1522
    hence "A = {}" by auto
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1523
    with empty show ?case by simp
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1524
  next
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1525
    case (Suc n A)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1526
    from \<open>A \<subseteq> B\<close> and \<open>finite B\<close> have "finite A" by (rule finite_subset)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1527
    moreover from Suc.hyps have "A \<noteq> {}" by auto
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1528
    moreover note \<open>A \<subseteq> B\<close>
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1529
    moreover have "P (A - {x})" if x: "x \<in> A" for x
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1530
      using x Suc.prems \<open>Suc n = card A\<close> by (intro Suc) auto
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1531
    ultimately show ?case by (rule remove)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1532
  qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1533
qed
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1534
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1535
lemma finite_remove_induct [consumes 1, case_names empty remove]:
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1536
  assumes finite: "finite B" and empty: "P ({} :: 'a set)" 
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  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"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1538
  defines "B' \<equiv> B"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1539
  shows   "P B'"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1540
  by (induction B' rule: remove_induct) (simp_all add: assms)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1541
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1542
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1543
text\<open>main cardinality theorem\<close>
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1544
lemma card_partition [rule_format]:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1545
  "finite C ==>
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60303
diff changeset
  1546
     finite (\<Union>C) -->
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1547
     (\<forall>c\<in>C. card c = k) -->
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1548
     (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 --> c1 \<inter> c2 = {}) -->
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60303
diff changeset
  1549
     k * card(C) = card (\<Union>C)"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1550
apply (erule finite_induct, simp)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1551
apply (simp add: card_Un_disjoint insert_partition 
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60303
diff changeset
  1552
       finite_subset [of _ "\<Union>(insert x F)"])
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1553
done
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1554
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1555
lemma card_eq_UNIV_imp_eq_UNIV:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1556
  assumes fin: "finite (UNIV :: 'a set)"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1557
  and card: "card A = card (UNIV :: 'a set)"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1558
  shows "A = (UNIV :: 'a set)"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1559
proof
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1560
  show "A \<subseteq> UNIV" by simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1561
  show "UNIV \<subseteq> A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1562
  proof
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1563
    fix x
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1564
    show "x \<in> A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1565
    proof (rule ccontr)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1566
      assume "x \<notin> A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1567
      then have "A \<subset> UNIV" by auto
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1568
      with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1569
      with card show False by simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1570
    qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1571
  qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1572
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1573
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1574
text\<open>The form of a finite set of given cardinality\<close>
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1575
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1576
lemma card_eq_SucD:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1577
assumes "card A = Suc k"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1578
shows "\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={})"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1579
proof -
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1580
  have fin: "finite A" using assms by (auto intro: ccontr)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1581
  moreover have "card A \<noteq> 0" using assms by auto
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1582
  ultimately obtain b where b: "b \<in> A" by auto
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1583
  show ?thesis
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1584
  proof (intro exI conjI)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1585
    show "A = insert b (A-{b})" using b by blast
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1586
    show "b \<notin> A - {b}" by blast
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1587
    show "card (A - {b}) = k" and "k = 0 \<longrightarrow> A - {b} = {}"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44835
diff changeset
  1588
      using assms b fin by(fastforce dest:mk_disjoint_insert)+
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1589
  qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1590
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1591
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1592
lemma card_Suc_eq:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1593
  "(card A = Suc k) =
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1594
   (\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={}))"
54570
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1595
 apply(auto elim!: card_eq_SucD)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1596
 apply(subst card.insert)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1597
 apply(auto simp add: intro:ccontr)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1598
 done
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1599
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61169
diff changeset
  1600
lemma card_1_singletonE:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61169
diff changeset
  1601
    assumes "card A = 1" obtains x where "A = {x}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61169
diff changeset
  1602
  using assms by (auto simp: card_Suc_eq)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61169
diff changeset
  1603
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1604
lemma is_singleton_altdef: "is_singleton A \<longleftrightarrow> card A = 1"
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1605
  unfolding is_singleton_def
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1606
  by (auto elim!: card_1_singletonE is_singletonE simp del: One_nat_def)
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 63040
diff changeset
  1607
44744
bdf8eb8f126b added new lemmas
nipkow
parents: 43991
diff changeset
  1608
lemma card_le_Suc_iff: "finite A \<Longrightarrow>
bdf8eb8f126b added new lemmas
nipkow
parents: 43991
diff changeset
  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)"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44835
diff changeset
  1610
by (fastforce simp: card_Suc_eq less_eq_nat.simps(2) insert_eq_iff
44744
bdf8eb8f126b added new lemmas
nipkow
parents: 43991
diff changeset
  1611
  dest: subset_singletonD split: nat.splits if_splits)
bdf8eb8f126b added new lemmas
nipkow
parents: 43991
diff changeset
  1612
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1613
lemma finite_fun_UNIVD2:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1614
  assumes fin: "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1615
  shows "finite (UNIV :: 'b set)"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1616
proof -
46146
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1617
  from fin have "\<And>arbitrary. finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary))"
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1618
    by (rule finite_imageI)
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1619
  moreover have "\<And>arbitrary. UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary)"
6baea4fca6bd incorporated various theorems from theory More_Set into corpus
haftmann
parents: 46033
diff changeset
  1620
    by (rule UNIV_eq_I) auto
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1621
  ultimately show "finite (UNIV :: 'b set)" by simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1622
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1623
48063
f02b4302d5dd remove duplicate lemma card_unit in favor of Finite_Set.card_UNIV_unit
huffman
parents: 47221
diff changeset
  1624
lemma card_UNIV_unit [simp]: "card (UNIV :: unit set) = 1"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1625
  unfolding UNIV_unit by simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1626
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1627
lemma infinite_arbitrarily_large:
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1628
  assumes "\<not> finite A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1629
  shows "\<exists>B. finite B \<and> card B = n \<and> B \<subseteq> A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1630
proof (induction n)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1631
  case 0 show ?case by (intro exI[of _ "{}"]) auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1632
next 
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1633
  case (Suc n)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1634
  then guess B .. note B = this
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1635
  with \<open>\<not> finite A\<close> have "A \<noteq> B" by auto
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1636
  with B have "B \<subset> A" by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1637
  hence "\<exists>x. x \<in> A - B" by (elim psubset_imp_ex_mem)
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1638
  then guess x .. note x = this
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1639
  with B have "finite (insert x B) \<and> card (insert x B) = Suc n \<and> insert x B \<subseteq> A"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1640
    by auto
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1641
  thus "\<exists>B. finite B \<and> card B = Suc n \<and> B \<subseteq> A" ..
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57025
diff changeset
  1642
qed
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1643
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1644
subsubsection \<open>Cardinality of image\<close>
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1645
54570
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1646
lemma card_image_le: "finite A ==> card (f ` A) \<le> card A"
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1647
  by (induct rule: finite_induct) (simp_all add: le_SucI card_insert_if)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1648
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1649
lemma card_image:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1650
  assumes "inj_on f A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1651
  shows "card (f ` A) = card A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1652
proof (cases "finite A")
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1653
  case True then show ?thesis using assms by (induct A) simp_all
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1654
next
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1655
  case False then have "\<not> finite (f ` A)" using assms by (auto dest: finite_imageD)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1656
  with False show ?thesis by simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1657
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1658
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1659
lemma bij_betw_same_card: "bij_betw f A B \<Longrightarrow> card A = card B"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1660
by(auto simp: card_image bij_betw_def)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1661
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1662
lemma endo_inj_surj: "finite A ==> f ` A \<subseteq> A ==> inj_on f A ==> f ` A = A"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1663
by (simp add: card_seteq card_image)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1664
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1665
lemma eq_card_imp_inj_on:
54570
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1666
  assumes "finite A" "card(f ` A) = card A" shows "inj_on f A"
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1667
using assms
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1668
proof (induct rule:finite_induct)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1669
  case empty show ?case by simp
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1670
next
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1671
  case (insert x A)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1672
  then show ?case using card_image_le [of A f]
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1673
    by (simp add: card_insert_if split: if_splits)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1674
qed
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1675
54570
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1676
lemma inj_on_iff_eq_card: "finite A \<Longrightarrow> inj_on f A \<longleftrightarrow> card(f ` A) = card A"
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1677
  by (blast intro: card_image eq_card_imp_inj_on)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1678
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1679
lemma card_inj_on_le:
54570
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1680
  assumes "inj_on f A" "f ` A \<subseteq> B" "finite B" shows "card A \<le> card B"
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1681
proof -
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1682
  have "finite A" using assms
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1683
    by (blast intro: finite_imageD dest: finite_subset)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1684
  then show ?thesis using assms 
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1685
   by (force intro: card_mono simp: card_image [symmetric])
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1686
qed
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1687
59504
8c6747dba731 New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents: 59336
diff changeset
  1688
lemma surj_card_le: "finite A \<Longrightarrow> B \<subseteq> f ` A \<Longrightarrow> card B \<le> card A"
8c6747dba731 New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents: 59336
diff changeset
  1689
  by (blast intro: card_image_le card_mono le_trans)
8c6747dba731 New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents: 59336
diff changeset
  1690
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1691
lemma card_bij_eq:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1692
  "[|inj_on f A; f ` A \<subseteq> B; inj_on g B; g ` B \<subseteq> A;
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1693
     finite A; finite B |] ==> card A = card B"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1694
by (auto intro: le_antisym card_inj_on_le)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1695
40703
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 40702
diff changeset
  1696
lemma bij_betw_finite:
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 40702
diff changeset
  1697
  assumes "bij_betw f A B"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 40702
diff changeset
  1698
  shows "finite A \<longleftrightarrow> finite B"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 40702
diff changeset
  1699
using assms unfolding bij_betw_def
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 40702
diff changeset
  1700
using finite_imageD[of f A] by auto
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1701
55020
96b05fd2aee4 dissolved 'Fun_More_FP' (a BNF dependency)
blanchet
parents: 54870
diff changeset
  1702
lemma inj_on_finite:
96b05fd2aee4 dissolved 'Fun_More_FP' (a BNF dependency)
blanchet
parents: 54870
diff changeset
  1703
assumes "inj_on f A" "f ` A \<le> B" "finite B"
96b05fd2aee4 dissolved 'Fun_More_FP' (a BNF dependency)
blanchet
parents: 54870
diff changeset
  1704
shows "finite A"
96b05fd2aee4 dissolved 'Fun_More_FP' (a BNF dependency)
blanchet
parents: 54870
diff changeset
  1705
using assms finite_imageD finite_subset by blast
96b05fd2aee4 dissolved 'Fun_More_FP' (a BNF dependency)
blanchet
parents: 54870
diff changeset
  1706
59520
76d7c593c6e8 add lema about card and vimage
Andreas Lochbihler
parents: 59519
diff changeset
  1707
lemma card_vimage_inj: "\<lbrakk> inj f; A \<subseteq> range f \<rbrakk> \<Longrightarrow> card (f -` A) = card A"
76d7c593c6e8 add lema about card and vimage
Andreas Lochbihler
parents: 59519
diff changeset
  1708
by(auto 4 3 simp add: subset_image_iff inj_vimage_image_eq intro: card_image[symmetric, OF subset_inj_on])
41656
011fcb70e32f restructured theory;
haftmann
parents: 41550
diff changeset
  1709
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1710
subsubsection \<open>Pigeonhole Principles\<close>
37466
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1711
40311
994e784ca17a removed assumption
nipkow
parents: 39302
diff changeset
  1712
lemma pigeonhole: "card A > card(f ` A) \<Longrightarrow> ~ inj_on f A "
37466
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1713
by (auto dest: card_image less_irrefl_nat)
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1714
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1715
lemma pigeonhole_infinite:
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1716
assumes  "~ finite A" and "finite(f`A)"
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1717
shows "EX a0:A. ~finite{a:A. f a = f a0}"
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1718
proof -
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1719
  have "finite(f`A) \<Longrightarrow> ~ finite A \<Longrightarrow> EX a0:A. ~finite{a:A. f a = f a0}"
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1720
  proof(induct "f`A" arbitrary: A rule: finite_induct)
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1721
    case empty thus ?case by simp
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1722
  next
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1723
    case (insert b F)
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1724
    show ?case
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1725
    proof cases
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1726
      assume "finite{a:A. f a = b}"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1727
      hence "~ finite(A - {a:A. f a = b})" using \<open>\<not> finite A\<close> by simp
37466
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1728
      also have "A - {a:A. f a = b} = {a:A. f a \<noteq> b}" by blast
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1729
      finally have "~ finite({a:A. f a \<noteq> b})" .
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1730
      from insert(3)[OF _ this]
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1731
      show ?thesis using insert(2,4) by simp (blast intro: rev_finite_subset)
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1732
    next
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1733
      assume 1: "~finite{a:A. f a = b}"
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1734
      hence "{a \<in> A. f a = b} \<noteq> {}" by force
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1735
      thus ?thesis using 1 by blast
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1736
    qed
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1737
  qed
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1738
  from this[OF assms(2,1)] show ?thesis .
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1739
qed
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1740
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1741
lemma pigeonhole_infinite_rel:
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1742
assumes "~finite A" and "finite B" and "ALL a:A. EX b:B. R a b"
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1743
shows "EX b:B. ~finite{a:A. R a b}"
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1744
proof -
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1745
   let ?F = "%a. {b:B. R a b}"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1746
   from finite_Pow_iff[THEN iffD2, OF \<open>finite B\<close>]
37466
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1747
   have "finite(?F ` A)" by(blast intro: rev_finite_subset)
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1748
   from pigeonhole_infinite[where f = ?F, OF assms(1) this]
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1749
   obtain a0 where "a0\<in>A" and 1: "\<not> finite {a\<in>A. ?F a = ?F a0}" ..
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1750
   obtain b0 where "b0 : B" and "R a0 b0" using \<open>a0:A\<close> assms(3) by blast
37466
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1751
   { assume "finite{a:A. R a b0}"
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1752
     then have "finite {a\<in>A. ?F a = ?F a0}"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1753
       using \<open>b0 : B\<close> \<open>R a0 b0\<close> by(blast intro: rev_finite_subset)
37466
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1754
   }
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1755
   with 1 \<open>b0 : B\<close> show ?thesis by blast
37466
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1756
qed
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1757
87bf104920f2 added pigeonhole lemmas
nipkow
parents: 36637
diff changeset
  1758
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1759
subsubsection \<open>Cardinality of sums\<close>
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1760
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1761
lemma card_Plus:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1762
  assumes "finite A" and "finite B"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1763
  shows "card (A <+> B) = card A + card B"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1764
proof -
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1765
  have "Inl`A \<inter> Inr`B = {}" by fast
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1766
  with assms show ?thesis
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1767
    unfolding Plus_def
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1768
    by (simp add: card_Un_disjoint card_image)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1769
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1770
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1771
lemma card_Plus_conv_if:
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1772
  "card (A <+> B) = (if finite A \<and> finite B then card A + card B else 0)"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1773
  by (auto simp add: card_Plus)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1774
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1775
text \<open>Relates to equivalence classes.  Based on a theorem of F. Kamm\"uller.\<close>
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1776
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1777
lemma dvd_partition:
54570
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  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 = {}"
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1779
    shows "k dvd card (\<Union>C)"
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1780
proof -
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1781
  have "finite C" 
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1782
    by (rule finite_UnionD [OF f])
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1783
  then show ?thesis using assms
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1784
  proof (induct rule: finite_induct)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1785
    case empty show ?case by simp
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1786
  next
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1787
    case (insert c C)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1788
    then show ?case 
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1789
      apply simp
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1790
      apply (subst card_Un_disjoint)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1791
      apply (auto simp add: disjoint_eq_subset_Compl)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1792
      done
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1793
  qed
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1794
qed
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1795
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60595
diff changeset
  1796
subsubsection \<open>Relating injectivity and surjectivity\<close>
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1797
54570
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1798
lemma finite_surj_inj: assumes "finite A" "A \<subseteq> f ` A" shows "inj_on f A"
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1799
proof -
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1800
  have "f ` A = A" 
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1801
    by (rule card_seteq [THEN sym]) (auto simp add: assms card_image_le)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1802
  then show ?thesis using assms
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1803
    by (simp add: eq_card_imp_inj_on)
002b8729f228 polished some ancient proofs
paulson
parents: 54413
diff changeset
  1804
qed
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1805
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1806
lemma finite_UNIV_surj_inj: fixes f :: "'a \<Rightarrow> 'a"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1807
shows "finite(UNIV:: 'a set) \<Longrightarrow> surj f \<Longrightarrow> inj f"
40702
cf26dd7395e4 Replace surj by abbreviation; remove surj_on.
hoelzl
parents: 40311
diff changeset
  1808
by (blast intro: finite_surj_inj subset_UNIV)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1809
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1810
lemma finite_UNIV_inj_surj: fixes f :: "'a \<Rightarrow> 'a"
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1811
shows "finite(UNIV:: 'a set) \<Longrightarrow> inj f \<Longrightarrow> surj f"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44835
diff changeset
  1812
by(fastforce simp:surj_def dest!: endo_inj_surj)
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1813
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1814
corollary infinite_UNIV_nat [iff]:
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1815
  "\<not> finite (UNIV :: nat set)"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1816
proof
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1817
  assume "finite (UNIV :: nat set)"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1818
  with finite_UNIV_inj_surj [of Suc]
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1819
  show False by simp (blast dest: Suc_neq_Zero surjD)
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1820
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1821
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53820
diff changeset
  1822
lemma infinite_UNIV_char_0:
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1823
  "\<not> finite (UNIV :: 'a::semiring_char_0 set)"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1824
proof
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1825
  assume "finite (UNIV :: 'a set)"
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1826
  with subset_UNIV have "finite (range of_nat :: 'a set)"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1827
    by (rule finite_subset)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1828
  moreover have "inj (of_nat :: nat \<Rightarrow> 'a)"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1829
    by (simp add: inj_on_def)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1830
  ultimately have "finite (UNIV :: nat set)"
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1831
    by (rule finite_imageD)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51487
diff changeset
  1832
  then show False
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1833
    by simp
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1834
qed
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1835
49758
718f10c8bbfc use Set.filter instead of Finite_Set.filter, which is removed then
kuncar
parents: 49757
diff changeset
  1836
hide_const (open) Finite_Set.fold
46033
6fc579c917b8 qualified Finite_Set.fold
haftmann
parents: 45962
diff changeset
  1837
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1838
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1839
subsection "Infinite Sets"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1840
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1841
text \<open>
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1842
  Some elementary facts about infinite sets, mostly by Stephan Merz.
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1843
  Beware! Because "infinite" merely abbreviates a negation, these
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1844
  lemmas may not work well with \<open>blast\<close>.
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1845
\<close>
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1846
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1847
abbreviation infinite :: "'a set \<Rightarrow> bool"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1848
  where "infinite S \<equiv> \<not> finite S"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1849
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1850
text \<open>
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1851
  Infinite sets are non-empty, and if we remove some elements from an
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1852
  infinite set, the result is still infinite.
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1853
\<close>
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1854
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1855
lemma infinite_imp_nonempty: "infinite S \<Longrightarrow> S \<noteq> {}"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1856
  by auto
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1857
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1858
lemma infinite_remove: "infinite S \<Longrightarrow> infinite (S - {a})"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1859
  by simp
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1860
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1861
lemma Diff_infinite_finite:
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1862
  assumes T: "finite T" and S: "infinite S"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1863
  shows "infinite (S - T)"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1864
  using T
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1865
proof induct
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1866
  from S
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1867
  show "infinite (S - {})" by auto
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1868
next
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1869
  fix T x
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1870
  assume ih: "infinite (S - T)"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1871
  have "S - (insert x T) = (S - T) - {x}"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1872
    by (rule Diff_insert)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1873
  with ih
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1874
  show "infinite (S - (insert x T))"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1875
    by (simp add: infinite_remove)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1876
qed
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1877
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1878
lemma Un_infinite: "infinite S \<Longrightarrow> infinite (S \<union> T)"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1879
  by simp
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1880
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1881
lemma infinite_Un: "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1882
  by simp
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1883
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1884
lemma infinite_super:
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1885
  assumes T: "S \<subseteq> T" and S: "infinite S"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1886
  shows "infinite T"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1887
proof
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1888
  assume "finite T"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1889
  with T have "finite S" by (simp add: finite_subset)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1890
  with S show False by simp
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1891
qed
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1892
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1893
proposition infinite_coinduct [consumes 1, case_names infinite]:
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1894
  assumes "X A"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1895
  and step: "\<And>A. X A \<Longrightarrow> \<exists>x\<in>A. X (A - {x}) \<or> infinite (A - {x})"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1896
  shows "infinite A"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1897
proof
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1898
  assume "finite A"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1899
  then show False using \<open>X A\<close>
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1900
  proof (induction rule: finite_psubset_induct)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1901
    case (psubset A)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1902
    then obtain x where "x \<in> A" "X (A - {x}) \<or> infinite (A - {x})"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1903
      using local.step psubset.prems by blast
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1904
    then have "X (A - {x})"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1905
      using psubset.hyps by blast
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1906
    show False
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1907
      apply (rule psubset.IH [where B = "A - {x}"])
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1908
      using \<open>x \<in> A\<close> apply blast
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1909
      by (simp add: \<open>X (A - {x})\<close>)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1910
  qed
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1911
qed
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1912
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1913
text \<open>
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1914
  For any function with infinite domain and finite range there is some
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1915
  element that is the image of infinitely many domain elements.  In
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1916
  particular, any infinite sequence of elements from a finite set
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1917
  contains some element that occurs infinitely often.
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1918
\<close>
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1919
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1920
lemma inf_img_fin_dom':
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1921
  assumes img: "finite (f ` A)" and dom: "infinite A"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1922
  shows "\<exists>y \<in> f ` A. infinite (f -` {y} \<inter> A)"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1923
proof (rule ccontr)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1924
  have "A \<subseteq> (\<Union>y\<in>f ` A. f -` {y} \<inter> A)" by auto
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1925
  moreover
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1926
  assume "\<not> ?thesis"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1927
  with img have "finite (\<Union>y\<in>f ` A. f -` {y} \<inter> A)" by blast
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1928
  ultimately have "finite A" by(rule finite_subset)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1929
  with dom show False by contradiction
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1930
qed
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1931
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1932
lemma inf_img_fin_domE':
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1933
  assumes "finite (f ` A)" and "infinite A"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1934
  obtains y where "y \<in> f`A" and "infinite (f -` {y} \<inter> A)"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1935
  using assms by (blast dest: inf_img_fin_dom')
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1936
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1937
lemma inf_img_fin_dom:
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1938
  assumes img: "finite (f`A)" and dom: "infinite A"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1939
  shows "\<exists>y \<in> f`A. infinite (f -` {y})"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1940
using inf_img_fin_dom'[OF assms] by auto
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1941
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1942
lemma inf_img_fin_domE:
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1943
  assumes "finite (f`A)" and "infinite A"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1944
  obtains y where "y \<in> f`A" and "infinite (f -` {y})"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1945
  using assms by (blast dest: inf_img_fin_dom)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1946
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1947
proposition finite_image_absD:
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1948
    fixes S :: "'a::linordered_ring set"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1949
    shows "finite (abs ` S) \<Longrightarrow> finite S"
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1950
  by (rule ccontr) (auto simp: abs_eq_iff vimage_def dest: inf_img_fin_dom)
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61799
diff changeset
  1951
35722
69419a09a7ff moved cardinality to Finite_Set as far as appropriate; added locales for fold_image
haftmann
parents: 35719
diff changeset
  1952
end