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