src/HOL/Finite_Set.thy
author kleing
Wed, 14 Apr 2004 14:13:05 +0200
changeset 14565 c6dc17aab88a
parent 14504 7a3d80e276d4
child 14661 9ead82084de8
permissions -rw-r--r--
use more symbols in HTML output
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
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
     4
                Additions by Jeremy Avigad in Feb 2004
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
     5
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
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
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
     8
header {* Finite sets *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
     9
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
    10
theory Finite_Set = Divides + Power + Inductive:
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    11
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    12
subsection {* Collection of finite sets *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    13
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    14
consts Finites :: "'a set set"
13737
e564c3d2d174 added a few lemmas
nipkow
parents: 13735
diff changeset
    15
syntax
e564c3d2d174 added a few lemmas
nipkow
parents: 13735
diff changeset
    16
  finite :: "'a set => bool"
e564c3d2d174 added a few lemmas
nipkow
parents: 13735
diff changeset
    17
translations
e564c3d2d174 added a few lemmas
nipkow
parents: 13735
diff changeset
    18
  "finite A" == "A : Finites"
12396
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
inductive Finites
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    21
  intros
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    22
    emptyI [simp, intro!]: "{} : Finites"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    23
    insertI [simp, intro!]: "A : Finites ==> insert 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
13737
e564c3d2d174 added a few lemmas
nipkow
parents: 13735
diff changeset
    28
lemma ex_new_if_finite: -- "does not depend on def of finite at all"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
    29
 "[| ~finite(UNIV::'a set); finite A  |] ==> \<exists>a::'a. a \<notin> A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
    30
by(subgoal_tac "A \<noteq> UNIV", blast, blast)
13737
e564c3d2d174 added a few lemmas
nipkow
parents: 13735
diff changeset
    31
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    32
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    33
lemma finite_induct [case_names empty insert, induct set: Finites]:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    34
  "finite F ==>
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    35
    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
    36
  -- {* Discharging @{text "x \<notin> F"} entails extra work. *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    37
proof -
13421
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
    38
  assume "P {}" and
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
    39
    insert: "!!F x. finite F ==> x \<notin> F ==> P F ==> P (insert x F)"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    40
  assume "finite F"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    41
  thus "P F"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    42
  proof induct
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    43
    show "P {}" .
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    44
    fix F x assume F: "finite F" and P: "P F"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    45
    show "P (insert x F)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    46
    proof cases
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    47
      assume "x \<in> F"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    48
      hence "insert x F = F" by (rule insert_absorb)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    49
      with P show ?thesis by (simp only:)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    50
    next
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    51
      assume "x \<notin> F"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    52
      from F this P show ?thesis by (rule insert)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    53
    qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    54
  qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    55
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    56
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    57
lemma finite_subset_induct [consumes 2, case_names empty insert]:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    58
  "finite F ==> F \<subseteq> A ==>
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    59
    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
    60
    P F"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    61
proof -
13421
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
    62
  assume "P {}" and insert:
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
    63
    "!!F a. finite F ==> a \<in> A ==> a \<notin> F ==> P F ==> P (insert a F)"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    64
  assume "finite F"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    65
  thus "F \<subseteq> A ==> P F"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    66
  proof induct
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    67
    show "P {}" .
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    68
    fix F x assume "finite F" and "x \<notin> F"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    69
      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
    70
    show "P (insert x F)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    71
    proof (rule insert)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    72
      from i show "x \<in> A" by blast
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    73
      from i have "F \<subseteq> A" by blast
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    74
      with P show "P F" .
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    75
    qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    76
  qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    77
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    78
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    79
lemma finite_UnI: "finite F ==> finite G ==> finite (F Un G)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    80
  -- {* The union of two finite sets is finite. *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    81
  by (induct set: Finites) simp_all
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    82
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    83
lemma finite_subset: "A \<subseteq> B ==> finite B ==> finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    84
  -- {* Every subset of a finite set is finite. *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    85
proof -
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    86
  assume "finite B"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    87
  thus "!!A. A \<subseteq> B ==> finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    88
  proof induct
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    89
    case empty
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    90
    thus ?case by simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    91
  next
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    92
    case (insert F x A)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    93
    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
    94
    show "finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    95
    proof cases
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    96
      assume x: "x \<in> A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    97
      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
    98
      with r have "finite (A - {x})" .
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
    99
      hence "finite (insert x (A - {x}))" ..
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   100
      also have "insert x (A - {x}) = A" by (rule insert_Diff)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   101
      finally show ?thesis .
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   102
    next
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   103
      show "A \<subseteq> F ==> ?thesis" .
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   104
      assume "x \<notin> A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   105
      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
   106
    qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   107
  qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   108
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   109
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   110
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
   111
  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
   112
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   113
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
   114
  -- {* The converse obviously fails. *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   115
  by (blast intro: finite_subset)
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_insert [simp]: "finite (insert a A) = finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   118
  apply (subst insert_is_Un)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   119
  apply (simp only: finite_Un, blast)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   120
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   121
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   122
lemma finite_empty_induct:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   123
  "finite A ==>
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   124
  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
   125
proof -
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   126
  assume "finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   127
    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
   128
  have "P (A - A)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   129
  proof -
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   130
    fix c b :: "'a set"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   131
    presume c: "finite c" and b: "finite b"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   132
      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
   133
    from c show "c \<subseteq> b ==> P (b - c)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   134
    proof induct
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   135
      case empty
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   136
      from P1 show ?case by simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   137
    next
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   138
      case (insert F x)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   139
      have "P (b - F - {x})"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   140
      proof (rule P2)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   141
        from _ b show "finite (b - F)" by (rule finite_subset) blast
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   142
        from insert show "x \<in> b - F" by simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   143
        from insert show "P (b - F)" by simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   144
      qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   145
      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
   146
      finally show ?case .
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   147
    qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   148
  next
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   149
    show "A \<subseteq> A" ..
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   150
  qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   151
  thus "P {}" by simp
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
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   154
lemma finite_Diff [simp]: "finite B ==> finite (B - Ba)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   155
  by (rule Diff_subset [THEN finite_subset])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   156
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   157
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
   158
  apply (subst Diff_insert)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   159
  apply (case_tac "a : A - B")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   160
   apply (rule finite_insert [symmetric, THEN trans])
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   161
   apply (subst insert_Diff, simp_all)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   162
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   163
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   164
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   165
subsubsection {* Image and Inverse Image over Finite Sets *}
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   166
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   167
lemma finite_imageI[simp]: "finite F ==> finite (h ` F)"
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   168
  -- {* The image of a finite set is finite. *}
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   169
  by (induct set: Finites) simp_all
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   170
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   171
lemma finite_surj: "finite A ==> B <= f ` A ==> finite B"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   172
  apply (frule finite_imageI)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   173
  apply (erule finite_subset, assumption)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   174
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   175
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   176
lemma finite_range_imageI:
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   177
    "finite (range g) ==> finite (range (%x. f (g x)))"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   178
  apply (drule finite_imageI, simp)
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   179
  done
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   180
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   181
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
   182
proof -
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   183
  have aux: "!!A. finite (A - {}) = finite A" by simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   184
  fix B :: "'a set"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   185
  assume "finite B"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   186
  thus "!!A. f`A = B ==> inj_on f A ==> finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   187
    apply induct
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   188
     apply simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   189
    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
   190
     apply clarify
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   191
     apply (simp (no_asm_use) add: inj_on_def)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   192
     apply (blast dest!: aux [THEN iffD1], atomize)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   193
    apply (erule_tac V = "ALL A. ?PP (A)" in thin_rl)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   194
    apply (frule subsetD [OF equalityD2 insertI1], clarify)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   195
    apply (rule_tac x = xa in bexI)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   196
     apply (simp_all add: inj_on_image_set_diff)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   197
    done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   198
qed (rule refl)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   199
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   200
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   201
lemma inj_vimage_singleton: "inj f ==> f-`{a} \<subseteq> {THE x. f x = a}"
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   202
  -- {* The inverse image of a singleton under an injective function
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   203
         is included in a singleton. *}
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   204
  apply (auto simp add: inj_on_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   205
  apply (blast intro: the_equality [symmetric])
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   206
  done
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   207
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   208
lemma finite_vimageI: "[|finite F; inj h|] ==> finite (h -` F)"
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   209
  -- {* The inverse image of a finite set under an injective function
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   210
         is finite. *}
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   211
  apply (induct set: Finites, simp_all)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   212
  apply (subst vimage_insert)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   213
  apply (simp add: finite_Un finite_subset [OF inj_vimage_singleton])
13825
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   214
  done
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   215
ef4c41e7956a new inverse image lemmas
paulson
parents: 13737
diff changeset
   216
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   217
subsubsection {* The finite UNION of finite sets *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   218
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   219
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
   220
  by (induct set: Finites) simp_all
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   221
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   222
text {*
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   223
  Strengthen RHS to
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   224
  @{prop "((ALL x:A. finite (B x)) & finite {x. x:A & B x \<noteq> {}})"}?
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   225
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   226
  We'd need to prove
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   227
  @{prop "finite C ==> ALL A B. (UNION A B) <= C --> finite {x. x:A & B x \<noteq> {}}"}
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   228
  by induction. *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   229
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   230
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
   231
  by (blast intro: finite_UN_I finite_subset)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   232
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   233
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   234
subsubsection {* Sigma of finite sets *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   235
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   236
lemma finite_SigmaI [simp]:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   237
    "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
   238
  by (unfold Sigma_def) (blast intro!: finite_UN_I)
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
lemma finite_Prod_UNIV:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   241
    "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
   242
  apply (subgoal_tac "(UNIV:: ('a * 'b) set) = Sigma UNIV (%x. UNIV)")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   243
   apply (erule ssubst)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   244
   apply (erule finite_SigmaI, auto)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   245
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   246
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   247
instance unit :: finite
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   248
proof
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   249
  have "finite {()}" by simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   250
  also have "{()} = UNIV" by auto
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   251
  finally show "finite (UNIV :: unit set)" .
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
instance * :: (finite, finite) finite
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   255
proof
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   256
  show "finite (UNIV :: ('a \<times> 'b) set)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   257
  proof (rule finite_Prod_UNIV)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   258
    show "finite (UNIV :: 'a set)" by (rule finite)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   259
    show "finite (UNIV :: 'b set)" by (rule finite)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   260
  qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   261
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   262
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   263
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   264
subsubsection {* The powerset of a finite set *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   265
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   266
lemma finite_Pow_iff [iff]: "finite (Pow A) = finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   267
proof
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   268
  assume "finite (Pow A)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   269
  with _ have "finite ((%x. {x}) ` A)" by (rule finite_subset) blast
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   270
  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
   271
next
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   272
  assume "finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   273
  thus "finite (Pow A)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   274
    by induct (simp_all add: finite_UnI finite_imageI Pow_insert)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   275
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   276
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   277
lemma finite_converse [iff]: "finite (r^-1) = finite r"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   278
  apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   279
   apply simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   280
   apply (rule iffI)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   281
    apply (erule finite_imageD [unfolded inj_on_def])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   282
    apply (simp split add: split_split)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   283
   apply (erule finite_imageI)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   284
  apply (simp add: converse_def image_def, auto)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   285
  apply (rule bexI)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   286
   prefer 2 apply assumption
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   287
  apply simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   288
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   289
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   290
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   291
subsubsection {* Finiteness of transitive closure *}
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
text {* (Thanks to Sidi Ehmety) *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   294
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   295
lemma finite_Field: "finite r ==> finite (Field r)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   296
  -- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   297
  apply (induct set: Finites)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   298
   apply (auto simp add: Field_def Domain_insert Range_insert)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   299
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   300
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   301
lemma trancl_subset_Field2: "r^+ <= Field r \<times> Field r"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   302
  apply clarify
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   303
  apply (erule trancl_induct)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   304
   apply (auto simp add: Field_def)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   305
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   306
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   307
lemma finite_trancl: "finite (r^+) = finite r"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   308
  apply auto
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   309
   prefer 2
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   310
   apply (rule trancl_subset_Field2 [THEN finite_subset])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   311
   apply (rule finite_SigmaI)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   312
    prefer 3
13704
854501b1e957 Transitive closure is now defined inductively as well.
berghofe
parents: 13595
diff changeset
   313
    apply (blast intro: r_into_trancl' finite_subset)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   314
   apply (auto simp add: finite_Field)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   315
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   316
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   317
lemma finite_cartesian_product: "[| finite A; finite B |] ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   318
    finite (A <*> B)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   319
  by (rule finite_SigmaI)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   320
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   321
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   322
subsection {* Finite cardinality *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   323
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   324
text {*
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   325
  This definition, although traditional, is ugly to work with: @{text
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   326
  "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
   327
  switched to an inductive one:
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
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   330
consts cardR :: "('a set \<times> nat) set"
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
inductive cardR
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   333
  intros
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   334
    EmptyI: "({}, 0) : cardR"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   335
    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
   336
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   337
constdefs
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   338
  card :: "'a set => nat"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   339
  "card A == THE n. (A, n) : cardR"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   340
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   341
inductive_cases cardR_emptyE: "({}, n) : cardR"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   342
inductive_cases cardR_insertE: "(insert a A,n) : cardR"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   343
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   344
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
   345
  by (induct set: cardR) simp_all
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   346
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   347
lemma cardR_determ_aux1:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   348
    "(A, m): cardR ==> (!!n a. m = Suc n ==> a:A ==> (A - {a}, n) : cardR)"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   349
  apply (induct set: cardR, auto)
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   350
  apply (simp add: insert_Diff_if, auto)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   351
  apply (drule cardR_SucD)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   352
  apply (blast intro!: cardR.intros)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   353
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   354
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   355
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
   356
  by (drule cardR_determ_aux1) auto
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   357
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   358
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
   359
  apply (induct set: cardR)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   360
   apply (safe elim!: cardR_emptyE cardR_insertE)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   361
  apply (rename_tac B b m)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   362
  apply (case_tac "a = b")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   363
   apply (subgoal_tac "A = B")
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   364
    prefer 2 apply (blast elim: equalityE, blast)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   365
  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
   366
   prefer 2
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   367
   apply (rule_tac x = "A Int B" in exI)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   368
   apply (blast elim: equalityE)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   369
  apply (frule_tac A = B in cardR_SucD)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   370
  apply (blast intro!: cardR.intros dest!: cardR_determ_aux2)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   371
  done
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 cardR_imp_finite: "(A, n) : cardR ==> finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   374
  by (induct set: cardR) simp_all
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 finite_imp_cardR: "finite A ==> EX n. (A, n) : cardR"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   377
  by (induct set: Finites) (auto intro!: cardR.intros)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   378
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   379
lemma card_equality: "(A,n) : cardR ==> card A = n"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   380
  by (unfold card_def) (blast intro: cardR_determ)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   381
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   382
lemma card_empty [simp]: "card {} = 0"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   383
  by (unfold card_def) (blast intro!: cardR.intros elim!: cardR_emptyE)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   384
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   385
lemma card_insert_disjoint [simp]:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   386
  "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
   387
proof -
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   388
  assume x: "x \<notin> A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   389
  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
   390
    apply (auto intro!: cardR.intros)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   391
    apply (rule_tac A1 = A in finite_imp_cardR [THEN exE])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   392
     apply (force dest: cardR_imp_finite)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   393
    apply (blast intro!: cardR.intros intro: cardR_determ)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   394
    done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   395
  assume "finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   396
  thus ?thesis
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   397
    apply (simp add: card_def aux)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   398
    apply (rule the_equality)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   399
     apply (auto intro: finite_imp_cardR
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   400
       cong: conj_cong simp: card_def [symmetric] card_equality)
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
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   403
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   404
lemma card_0_eq [simp]: "finite A ==> (card A = 0) = (A = {})"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   405
  apply auto
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   406
  apply (drule_tac a = x in mk_disjoint_insert, clarify)
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   407
  apply (rotate_tac -1, auto)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   408
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   409
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   410
lemma card_insert_if:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   411
    "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
   412
  by (simp add: insert_absorb)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   413
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   414
lemma card_Suc_Diff1: "finite A ==> x: A ==> Suc (card (A - {x})) = card A"
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
   415
apply(rule_tac t = A in insert_Diff [THEN subst], assumption)
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
   416
apply(simp del:insert_Diff_single)
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
   417
done
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   418
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   419
lemma card_Diff_singleton:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   420
    "finite A ==> x: A ==> card (A - {x}) = card A - 1"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   421
  by (simp add: card_Suc_Diff1 [symmetric])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   422
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   423
lemma card_Diff_singleton_if:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   424
    "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
   425
  by (simp add: card_Diff_singleton)
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_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   428
  by (simp add: card_insert_if card_Suc_Diff1)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   429
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   430
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
   431
  by (simp add: card_insert_if)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   432
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   433
lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   434
  apply (induct set: Finites, simp, clarify)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   435
  apply (subgoal_tac "finite A & A - {x} <= F")
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   436
   prefer 2 apply (blast intro: finite_subset, atomize)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   437
  apply (drule_tac x = "A - {x}" in spec)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   438
  apply (simp add: card_Diff_singleton_if split add: split_if_asm)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   439
  apply (case_tac "card A", auto)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   440
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   441
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   442
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
   443
  apply (simp add: psubset_def linorder_not_le [symmetric])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   444
  apply (blast dest: card_seteq)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   445
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   446
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   447
lemma card_mono: "finite B ==> A <= B ==> card A <= card B"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   448
  apply (case_tac "A = B", simp)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   449
  apply (simp add: linorder_not_less [symmetric])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   450
  apply (blast dest: card_seteq intro: order_less_imp_le)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   451
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   452
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   453
lemma card_Un_Int: "finite A ==> finite B
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   454
    ==> card A + card B = card (A Un B) + card (A Int B)"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   455
  apply (induct set: Finites, simp)
12396
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
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14302
diff changeset
   467
  apply (rule nat_add_right_cancel [THEN iffD1])
12396
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"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   494
by (erule psubsetI, blast)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   495
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   496
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   497
subsubsection {* Cardinality of image *}
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
lemma card_image_le: "finite A ==> card (f ` A) <= card A"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   500
  apply (induct set: Finites, simp)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   501
  apply (simp add: le_SucI finite_imageI card_insert_if)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   502
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   503
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   504
lemma card_image: "finite A ==> inj_on f A ==> card (f ` A) = card A"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   505
  apply (induct set: Finites, simp_all, atomize, safe)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   506
   apply (unfold inj_on_def, blast)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   507
  apply (subst card_insert_disjoint)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   508
    apply (erule finite_imageI, blast, blast)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   509
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   510
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   511
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
   512
  by (simp add: card_seteq card_image)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   513
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   514
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   515
subsubsection {* Cardinality of the Powerset *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   516
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   517
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
   518
  apply (induct set: Finites)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   519
   apply (simp_all add: Pow_insert)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   520
  apply (subst card_Un_disjoint, blast)
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   521
    apply (blast intro: finite_imageI, blast)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   522
  apply (subgoal_tac "inj_on (insert x) (Pow F)")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   523
   apply (simp add: card_image Pow_insert)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   524
  apply (unfold inj_on_def)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   525
  apply (blast elim!: equalityE)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   526
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   527
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   528
text {*
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   529
  \medskip Relates to equivalence classes.  Based on a theorem of
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   530
  F. Kammüller's.  The @{prop "finite C"} premise is redundant.
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   531
*}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   532
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   533
lemma dvd_partition:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   534
  "finite C ==> finite (Union C) ==>
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   535
    ALL c : C. k dvd card c ==>
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   536
    (ALL c1: C. ALL c2: C. c1 \<noteq> c2 --> c1 Int c2 = {}) ==>
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   537
  k dvd card (Union C)"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   538
  apply (induct set: Finites, simp_all, clarify)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   539
  apply (subst card_Un_disjoint)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   540
  apply (auto simp add: dvd_add disjoint_eq_subset_Compl)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   541
  done
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
subsection {* A fold functional for finite sets *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   545
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   546
text {*
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   547
  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
   548
  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
   549
*}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   550
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   551
consts
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   552
  foldSet :: "('b => 'a => 'a) => 'a => ('b set \<times> 'a) set"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   553
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   554
inductive "foldSet f e"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   555
  intros
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   556
    emptyI [intro]: "({}, e) : foldSet f e"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   557
    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
   558
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   559
inductive_cases empty_foldSetE [elim!]: "({}, x) : foldSet f e"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   560
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   561
constdefs
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   562
  fold :: "('b => 'a => 'a) => 'a => 'b set => 'a"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   563
  "fold f e A == THE x. (A, x) : foldSet f e"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   564
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   565
lemma Diff1_foldSet: "(A - {x}, y) : foldSet f e ==> x: A ==> (A, f x y) : foldSet f e"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   566
by (erule insert_Diff [THEN subst], rule foldSet.intros, auto)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   567
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   568
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
   569
  by (induct set: foldSet) auto
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   570
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   571
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
   572
  by (induct set: Finites) auto
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
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   575
subsubsection {* Left-commutative operations *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   576
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   577
locale LC =
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   578
  fixes f :: "'b => 'a => 'a"    (infixl "\<cdot>" 70)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   579
  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
   580
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   581
lemma (in LC) foldSet_determ_aux:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   582
  "ALL A x. card A < n --> (A, x) : foldSet f e -->
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   583
    (ALL y. (A, y) : foldSet f e --> y = x)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   584
  apply (induct n)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   585
   apply (auto simp add: less_Suc_eq)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   586
  apply (erule foldSet.cases, blast)
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   587
  apply (erule foldSet.cases, blast, clarify)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   588
  txt {* force simplification of @{text "card A < card (insert ...)"}. *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   589
  apply (erule rev_mp)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   590
  apply (simp add: less_Suc_eq_le)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   591
  apply (rule impI)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   592
  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
   593
   apply (subgoal_tac "Aa = Ab")
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   594
    prefer 2 apply (blast elim!: equalityE, blast)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   595
  txt {* case @{prop "xa \<notin> xb"}. *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   596
  apply (subgoal_tac "Aa - {xb} = Ab - {xa} & xb : Aa & xa : Ab")
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   597
   prefer 2 apply (blast elim!: equalityE, clarify)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   598
  apply (subgoal_tac "Aa = insert xb Ab - {xa}")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   599
   prefer 2 apply blast
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   600
  apply (subgoal_tac "card Aa <= card Ab")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   601
   prefer 2
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   602
   apply (rule Suc_le_mono [THEN subst])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   603
   apply (simp add: card_Suc_Diff1)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   604
  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
   605
  apply (blast intro: foldSet_imp_finite finite_Diff)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   606
  apply (frule (1) Diff1_foldSet)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   607
  apply (subgoal_tac "ya = f xb x")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   608
   prefer 2 apply (blast del: equalityCE)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   609
  apply (subgoal_tac "(Ab - {xa}, x) : foldSet f e")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   610
   prefer 2 apply simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   611
  apply (subgoal_tac "yb = f xa x")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   612
   prefer 2 apply (blast del: equalityCE dest: Diff1_foldSet)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   613
  apply (simp (no_asm_simp) add: left_commute)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   614
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   615
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   616
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
   617
  by (blast intro: foldSet_determ_aux [rule_format])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   618
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   619
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
   620
  by (unfold fold_def) (blast intro: foldSet_determ)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   621
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   622
lemma fold_empty [simp]: "fold f e {} = e"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   623
  by (unfold fold_def) blast
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   624
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   625
lemma (in LC) fold_insert_aux: "x \<notin> A ==>
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   626
    ((insert x A, v) : foldSet f e) =
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   627
    (EX y. (A, y) : foldSet f e & v = f x y)"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   628
  apply auto
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   629
  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
   630
   apply (fastsimp dest: foldSet_imp_finite)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   631
  apply (blast intro: foldSet_determ)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   632
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   633
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   634
lemma (in LC) fold_insert:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   635
    "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
   636
  apply (unfold fold_def)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   637
  apply (simp add: fold_insert_aux)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   638
  apply (rule the_equality)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   639
  apply (auto intro: finite_imp_foldSet
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   640
    cong add: conj_cong simp add: fold_def [symmetric] fold_equality)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   641
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   642
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   643
lemma (in LC) fold_commute: "finite A ==> (!!e. f x (fold f e A) = fold f (f x e) A)"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   644
  apply (induct set: Finites, simp)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   645
  apply (simp add: left_commute fold_insert)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   646
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   647
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   648
lemma (in LC) fold_nest_Un_Int:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   649
  "finite A ==> finite B
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   650
    ==> fold f (fold f e B) A = fold f (fold f e (A Int B)) (A Un B)"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   651
  apply (induct set: Finites, simp)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   652
  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
   653
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   654
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   655
lemma (in LC) fold_nest_Un_disjoint:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   656
  "finite A ==> finite B ==> A Int B = {}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   657
    ==> 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
   658
  by (simp add: fold_nest_Un_Int)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   659
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   660
declare foldSet_imp_finite [simp del]
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   661
    empty_foldSetE [rule del]  foldSet.intros [rule del]
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   662
  -- {* Delete rules to do with @{text foldSet} relation. *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   663
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   664
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   665
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   666
subsubsection {* Commutative monoids *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   667
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   668
text {*
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   669
  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
   670
  instead of @{text "'b => 'a => 'a"}.
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   671
*}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   672
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   673
locale ACe =
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   674
  fixes f :: "'a => 'a => 'a"    (infixl "\<cdot>" 70)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   675
    and e :: 'a
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   676
  assumes ident [simp]: "x \<cdot> e = x"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   677
    and commute: "x \<cdot> y = y \<cdot> x"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   678
    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
   679
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   680
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
   681
proof -
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   682
  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
   683
  also have "... = y \<cdot> (z \<cdot> x)" by (simp only: assoc)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   684
  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
   685
  finally show ?thesis .
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   686
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   687
12718
ade42a6c22ad lemmas (in ACe) AC;
wenzelm
parents: 12693
diff changeset
   688
lemmas (in ACe) AC = assoc commute left_commute
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   689
12693
827818b891c7 qualified exports from locales;
wenzelm
parents: 12396
diff changeset
   690
lemma (in ACe) left_ident [simp]: "e \<cdot> x = x"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   691
proof -
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   692
  have "x \<cdot> e = x" by (rule ident)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   693
  thus ?thesis by (subst commute)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   694
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   695
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   696
lemma (in ACe) fold_Un_Int:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   697
  "finite A ==> finite B ==>
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   698
    fold f e A \<cdot> fold f e B = fold f e (A Un B) \<cdot> fold f e (A Int B)"
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   699
  apply (induct set: Finites, simp)
13400
dbb608cd11c4 accomodate cumulative locale predicates;
wenzelm
parents: 13390
diff changeset
   700
  apply (simp add: AC insert_absorb Int_insert_left
13421
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
   701
    LC.fold_insert [OF LC.intro])
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   702
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   703
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   704
lemma (in ACe) fold_Un_disjoint:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   705
  "finite A ==> finite B ==> A Int B = {} ==>
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   706
    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
   707
  by (simp add: fold_Un_Int)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   708
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   709
lemma (in ACe) fold_Un_disjoint2:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   710
  "finite A ==> finite B ==> A Int B = {} ==>
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   711
    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
   712
proof -
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   713
  assume b: "finite B"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   714
  assume "finite A"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   715
  thus "A Int B = {} ==>
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   716
    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
   717
  proof induct
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   718
    case empty
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   719
    thus ?case by simp
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   720
  next
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   721
    case (insert F x)
13571
d76a798281f4 less use of x-symbols
paulson
parents: 13490
diff changeset
   722
    have "fold (f o g) e (insert x F \<union> B) = fold (f o g) e (insert x (F \<union> B))"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   723
      by simp
13571
d76a798281f4 less use of x-symbols
paulson
parents: 13490
diff changeset
   724
    also have "... = (f o g) x (fold (f o g) e (F \<union> B))"
13400
dbb608cd11c4 accomodate cumulative locale predicates;
wenzelm
parents: 13390
diff changeset
   725
      by (rule LC.fold_insert [OF LC.intro])
13421
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
   726
        (insert b insert, auto simp add: left_commute)
13571
d76a798281f4 less use of x-symbols
paulson
parents: 13490
diff changeset
   727
    also from insert have "fold (f o g) e (F \<union> B) =
d76a798281f4 less use of x-symbols
paulson
parents: 13490
diff changeset
   728
      fold (f o g) e F \<cdot> fold (f o g) e B" by blast
d76a798281f4 less use of x-symbols
paulson
parents: 13490
diff changeset
   729
    also have "(f o g) x ... = (f o g) x (fold (f o g) e F) \<cdot> fold (f o g) e B"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   730
      by (simp add: AC)
13571
d76a798281f4 less use of x-symbols
paulson
parents: 13490
diff changeset
   731
    also have "(f o g) x (fold (f o g) e F) = fold (f o g) e (insert x F)"
13400
dbb608cd11c4 accomodate cumulative locale predicates;
wenzelm
parents: 13390
diff changeset
   732
      by (rule LC.fold_insert [OF LC.intro, symmetric]) (insert b insert,
13421
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
   733
	auto simp add: left_commute)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   734
    finally show ?case .
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   735
  qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   736
qed
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   737
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   738
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   739
subsection {* Generalized summation over a set *}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   740
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   741
constdefs
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   742
  setsum :: "('a => 'b) => 'a set => 'b::plus_ac0"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   743
  "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
   744
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   745
syntax
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   746
  "_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
   747
syntax (xsymbols)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   748
  "_setsum" :: "idt => 'a set => 'b => 'b::plus_ac0"    ("\<Sum>_\<in>_. _" [0, 51, 10] 10)
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14504
diff changeset
   749
syntax (HTML output)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14504
diff changeset
   750
  "_setsum" :: "idt => 'a set => 'b => 'b::plus_ac0"    ("\<Sum>_\<in>_. _" [0, 51, 10] 10)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   751
translations
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   752
  "\<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
   753
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   754
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   755
lemma setsum_empty [simp]: "setsum f {} = 0"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   756
  by (simp add: setsum_def)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   757
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   758
lemma setsum_insert [simp]:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   759
    "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"
13365
a2c4faad4d35 adapted to locale defs;
wenzelm
parents: 12937
diff changeset
   760
  by (simp add: setsum_def
13421
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
   761
    LC.fold_insert [OF LC.intro] plus_ac0_left_commute)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   762
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   763
lemma setsum_reindex [rule_format]: "finite B ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   764
                  inj_on f B --> setsum h (f ` B) = setsum (h \<circ> f) B"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   765
apply (rule finite_induct, assumption, force)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   766
apply (rule impI, auto)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   767
apply (simp add: inj_on_def)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   768
apply (subgoal_tac "f x \<notin> f ` F")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   769
apply (subgoal_tac "finite (f ` F)")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   770
apply (auto simp add: setsum_insert)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   771
apply (simp add: inj_on_def)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   772
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   773
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   774
lemma setsum_reindex_id: "finite B ==> inj_on f B ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   775
    setsum f B = setsum id (f ` B)"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   776
by (auto simp add: setsum_reindex id_o)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   777
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   778
lemma setsum_reindex_cong: "finite A ==> inj_on f A ==> 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   779
    B = f ` A ==> g = h \<circ> f ==> setsum h B = setsum g A"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   780
  by (frule setsum_reindex, assumption, simp)
12396
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_cong:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   783
  "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
   784
  apply (case_tac "finite B")
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   785
   prefer 2 apply (simp add: setsum_def, simp)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   786
  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
   787
   apply simp
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   788
  apply (erule finite_induct, simp)
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
   789
  apply (simp add: subset_insert_iff, clarify)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   790
  apply (subgoal_tac "finite C")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   791
   prefer 2 apply (blast dest: finite_subset [COMP swap_prems_rl])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   792
  apply (subgoal_tac "C = insert x (C - {x})")
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   793
   prefer 2 apply blast
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   794
  apply (erule ssubst)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   795
  apply (drule spec)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   796
  apply (erule (1) notE impE)
14302
6c24235e8d5d *** empty log message ***
nipkow
parents: 14208
diff changeset
   797
  apply (simp add: Ball_def del:insert_Diff_single)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   798
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
   799
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   800
lemma setsum_0: "setsum (%i. 0) A = 0"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   801
  apply (case_tac "finite A")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   802
   prefer 2 apply (simp add: setsum_def)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   803
  apply (erule finite_induct, auto)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   804
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   805
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   806
lemma setsum_0': "ALL a:F. f a = 0 ==> setsum f F = 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   807
  apply (subgoal_tac "setsum f F = setsum (%x. 0) F")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   808
  apply (erule ssubst, rule setsum_0)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   809
  apply (rule setsum_cong, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   810
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   811
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   812
lemma card_eq_setsum: "finite A ==> card A = setsum (%x. 1) A"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   813
  -- {* Could allow many @{text "card"} proofs to be simplified. *}
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   814
  by (induct set: Finites) auto
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   815
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   816
lemma setsum_Un_Int: "finite A ==> finite B
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   817
    ==> setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   818
  -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   819
  apply (induct set: Finites, simp)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   820
  apply (simp add: plus_ac0 Int_insert_left insert_absorb)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   821
  done
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   822
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   823
lemma setsum_Un_disjoint: "finite A ==> finite B
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   824
  ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   825
  apply (subst setsum_Un_Int [symmetric], auto)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   826
  done
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   827
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   828
lemma setsum_UN_disjoint:
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   829
    "finite I ==> (ALL i:I. finite (A i)) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   830
        (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   831
      setsum f (UNION I A) = setsum (%i. setsum f (A i)) I"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   832
  apply (induct set: Finites, simp, atomize)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   833
  apply (subgoal_tac "ALL i:F. x \<noteq> i")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   834
   prefer 2 apply blast
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   835
  apply (subgoal_tac "A x Int UNION F A = {}")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   836
   prefer 2 apply blast
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   837
  apply (simp add: setsum_Un_disjoint)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   838
  done
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   839
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   840
lemma setsum_Union_disjoint:
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   841
  "finite C ==> (ALL A:C. finite A) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   842
        (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   843
      setsum f (Union C) = setsum (setsum f) C"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   844
  apply (frule setsum_UN_disjoint [of C id f])
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   845
  apply (unfold Union_def id_def, assumption+)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   846
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   847
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   848
lemma setsum_Sigma: "finite A ==> ALL x:A. finite (B x) ==> 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   849
    (\<Sum> x:A. (\<Sum> y:B x. f x y)) = 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   850
    (\<Sum> z:(SIGMA x:A. B x). f (fst z) (snd z))"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   851
  apply (subst Sigma_def)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   852
  apply (subst setsum_UN_disjoint)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   853
  apply assumption
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   854
  apply (rule ballI)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   855
  apply (drule_tac x = i in bspec, assumption)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   856
  apply (subgoal_tac "(UN y:(B i). {(i, y)}) <= (%y. (i, y)) ` (B i)") 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   857
  apply (rule finite_surj)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   858
  apply auto
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   859
  apply (rule setsum_cong, rule refl)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   860
  apply (subst setsum_UN_disjoint)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   861
  apply (erule bspec, assumption)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   862
  apply auto
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   863
  done
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   864
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   865
lemma setsum_cartesian_product: "finite A ==> finite B ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   866
    (\<Sum> x:A. (\<Sum> y:B. f x y)) = 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   867
    (\<Sum> z:A <*> B. f (fst z) (snd z))"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   868
  by (erule setsum_Sigma, auto);
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   869
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   870
lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   871
  apply (case_tac "finite A")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   872
   prefer 2 apply (simp add: setsum_def)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   873
  apply (erule finite_induct, auto)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   874
  apply (simp add: plus_ac0)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   875
  done
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   876
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   877
subsubsection {* Properties in more restricted classes of structures *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   878
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   879
lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   880
  apply (case_tac "finite A")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   881
   prefer 2 apply (simp add: setsum_def)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   882
  apply (erule rev_mp)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   883
  apply (erule finite_induct, auto)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   884
  done
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   885
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   886
lemma setsum_eq_0_iff [simp]:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   887
    "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   888
  by (induct set: Finites) auto
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   889
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   890
lemma setsum_constant_nat [simp]:
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   891
    "finite A ==> (\<Sum>x: A. y) = (card A) * y"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   892
  -- {* Later generalized to any semiring. *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   893
  by (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   894
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   895
lemma setsum_Un: "finite A ==> finite B ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   896
    (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   897
  -- {* For the natural numbers, we have subtraction. *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   898
  by (subst setsum_Un_Int [symmetric], auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   899
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   900
lemma setsum_Un_ring: "finite A ==> finite B ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   901
    (setsum f (A Un B) :: 'a :: ring) =
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   902
      setsum f A + setsum f B - setsum f (A Int B)"
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14485
diff changeset
   903
  by (subst setsum_Un_Int [symmetric], auto)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   904
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   905
lemma setsum_diff1: "(setsum f (A - {a}) :: nat) =
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   906
    (if a:A then setsum f A - f a else setsum f A)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   907
  apply (case_tac "finite A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   908
   prefer 2 apply (simp add: setsum_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   909
  apply (erule finite_induct)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   910
   apply (auto simp add: insert_Diff_if)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   911
  apply (drule_tac a = a in mk_disjoint_insert, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   912
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   913
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   914
lemma setsum_negf: "finite A ==> setsum (%x. - (f x)::'a::ring) A =
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   915
  - setsum f A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   916
  by (induct set: Finites, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   917
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   918
lemma setsum_subtractf: "finite A ==> setsum (%x. ((f x)::'a::ring) - g x) A =
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   919
  setsum f A - setsum g A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   920
  by (simp add: diff_minus setsum_addf setsum_negf)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   921
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   922
lemma setsum_nonneg: "[| finite A;
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   923
    \<forall>x \<in> A. (0::'a::ordered_semiring) \<le> f x |] ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   924
    0 \<le>  setsum f A";
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   925
  apply (induct set: Finites, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   926
  apply (subgoal_tac "0 + 0 \<le> f x + setsum f F", simp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   927
  apply (blast intro: add_mono)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   928
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   929
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   930
subsubsection {* Cardinality of unions and Sigma sets *}
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   931
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   932
lemma card_UN_disjoint:
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   933
    "finite I ==> (ALL i:I. finite (A i)) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   934
        (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   935
      card (UNION I A) = setsum (%i. card (A i)) I"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   936
  apply (subst card_eq_setsum)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   937
  apply (subst finite_UN, assumption+)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   938
  apply (subgoal_tac "setsum (%i. card (A i)) I =
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   939
      setsum (%i. (setsum (%x. 1) (A i))) I")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   940
  apply (erule ssubst)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   941
  apply (erule setsum_UN_disjoint, assumption+)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   942
  apply (rule setsum_cong)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   943
  apply simp+
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   944
  done
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   945
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   946
lemma card_Union_disjoint:
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   947
  "finite C ==> (ALL A:C. finite A) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   948
        (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   949
      card (Union C) = setsum card C"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   950
  apply (frule card_UN_disjoint [of C id])
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   951
  apply (unfold Union_def id_def, assumption+)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   952
  done
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   953
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   954
lemma SigmaI_insert: "y \<notin> A ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   955
  (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \<union> (SIGMA x: A. B x))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   956
  by auto
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   957
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   958
lemma card_cartesian_product_singleton: "finite A ==>
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   959
    card({x} <*> A) = card(A)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   960
  apply (subgoal_tac "inj_on (%y .(x,y)) A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   961
  apply (frule card_image, assumption)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   962
  apply (subgoal_tac "(Pair x ` A) = {x} <*> A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   963
  apply (auto simp add: inj_on_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   964
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   965
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   966
lemma card_SigmaI [rule_format,simp]: "finite A ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   967
  (ALL a:A. finite (B a)) -->
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   968
  card (SIGMA x: A. B x) = (\<Sum>a: A. card (B a))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   969
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   970
  apply (subst SigmaI_insert, assumption)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   971
  apply (subst card_Un_disjoint)
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   972
  apply (auto intro: finite_SigmaI simp add: card_cartesian_product_singleton)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   973
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   974
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   975
lemma card_cartesian_product: "[| finite A; finite B |] ==>
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   976
  card (A <*> B) = card(A) * card(B)"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
   977
  by simp
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   978
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   979
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   980
subsection {* Generalized product over a set *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   981
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   982
constdefs
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   983
  setprod :: "('a => 'b) => 'a set => 'b::times_ac1"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   984
  "setprod f A == if finite A then fold (op * o f) 1 A else 1"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   985
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   986
syntax
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   987
  "_setprod" :: "idt => 'a set => 'b => 'b::plus_ac0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   988
                ("\<Prod>_:_. _" [0, 51, 10] 10)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   989
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   990
syntax (xsymbols)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   991
  "_setprod" :: "idt => 'a set => 'b => 'b::plus_ac0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   992
                ("\<Prod>_\<in>_. _" [0, 51, 10] 10)
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14504
diff changeset
   993
syntax (HTML output)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14504
diff changeset
   994
  "_setprod" :: "idt => 'a set => 'b => 'b::plus_ac0"
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14504
diff changeset
   995
                ("\<Prod>_\<in>_. _" [0, 51, 10] 10)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   996
translations
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   997
  "\<Prod>i:A. b" == "setprod (%i. b) A"  -- {* Beware of argument permutation! *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   998
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
   999
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1000
lemma setprod_empty [simp]: "setprod f {} = 1"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1001
  by (auto simp add: setprod_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1002
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1003
lemma setprod_insert [simp]: "[| finite A; a \<notin> A |] ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1004
    setprod f (insert a A) = f a * setprod f A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1005
  by (auto simp add: setprod_def LC_def LC.fold_insert
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1006
      times_ac1_left_commute)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1007
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1008
lemma setprod_reindex [rule_format]: "finite B ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1009
                  inj_on f B --> setprod h (f ` B) = setprod (h \<circ> f) B"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1010
apply (rule finite_induct, assumption, force)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1011
apply (rule impI, auto)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1012
apply (simp add: inj_on_def)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1013
apply (subgoal_tac "f x \<notin> f ` F")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1014
apply (subgoal_tac "finite (f ` F)")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1015
apply (auto simp add: setprod_insert)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1016
apply (simp add: inj_on_def)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1017
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1018
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1019
lemma setprod_reindex_id: "finite B ==> inj_on f B ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1020
    setprod f B = setprod id (f ` B)"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1021
by (auto simp add: setprod_reindex id_o)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1022
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1023
lemma setprod_reindex_cong: "finite A ==> inj_on f A ==> 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1024
    B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1025
  by (frule setprod_reindex, assumption, simp)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1026
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1027
lemma setprod_cong:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1028
  "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1029
  apply (case_tac "finite B")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1030
   prefer 2 apply (simp add: setprod_def, simp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1031
  apply (subgoal_tac "ALL C. C <= B --> (ALL x:C. f x = g x) --> setprod f C = setprod g C")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1032
   apply simp
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1033
  apply (erule finite_induct, simp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1034
  apply (simp add: subset_insert_iff, clarify)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1035
  apply (subgoal_tac "finite C")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1036
   prefer 2 apply (blast dest: finite_subset [COMP swap_prems_rl])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1037
  apply (subgoal_tac "C = insert x (C - {x})")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1038
   prefer 2 apply blast
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1039
  apply (erule ssubst)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1040
  apply (drule spec)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1041
  apply (erule (1) notE impE)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1042
  apply (simp add: Ball_def del:insert_Diff_single)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1043
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1044
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1045
lemma setprod_1: "setprod (%i. 1) A = 1"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1046
  apply (case_tac "finite A")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1047
  apply (erule finite_induct, auto simp add: times_ac1)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1048
  apply (simp add: setprod_def)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1049
  done
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1050
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1051
lemma setprod_1': "ALL a:F. f a = 1 ==> setprod f F = 1"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1052
  apply (subgoal_tac "setprod f F = setprod (%x. 1) F")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1053
  apply (erule ssubst, rule setprod_1)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1054
  apply (rule setprod_cong, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1055
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1056
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1057
lemma setprod_Un_Int: "finite A ==> finite B
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1058
    ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1059
  apply (induct set: Finites, simp)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1060
  apply (simp add: times_ac1 insert_absorb)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1061
  apply (simp add: times_ac1 Int_insert_left insert_absorb)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1062
  done
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1063
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1064
lemma setprod_Un_disjoint: "finite A ==> finite B
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1065
  ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1066
  apply (subst setprod_Un_Int [symmetric], auto simp add: times_ac1)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1067
  done
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1068
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1069
lemma setprod_UN_disjoint:
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1070
    "finite I ==> (ALL i:I. finite (A i)) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1071
        (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1072
      setprod f (UNION I A) = setprod (%i. setprod f (A i)) I"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1073
  apply (induct set: Finites, simp, atomize)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1074
  apply (subgoal_tac "ALL i:F. x \<noteq> i")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1075
   prefer 2 apply blast
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1076
  apply (subgoal_tac "A x Int UNION F A = {}")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1077
   prefer 2 apply blast
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1078
  apply (simp add: setprod_Un_disjoint)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1079
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1080
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1081
lemma setprod_Union_disjoint:
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1082
  "finite C ==> (ALL A:C. finite A) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1083
        (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) ==>
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1084
      setprod f (Union C) = setprod (setprod f) C"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1085
  apply (frule setprod_UN_disjoint [of C id f])
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1086
  apply (unfold Union_def id_def, assumption+)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1087
  done
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1088
14485
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1089
lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==> 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1090
    (\<Prod> x:A. (\<Prod> y: B x. f x y)) = 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1091
    (\<Prod> z:(SIGMA x:A. B x). f (fst z) (snd z))"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1092
  apply (subst Sigma_def)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1093
  apply (subst setprod_UN_disjoint)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1094
  apply assumption
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1095
  apply (rule ballI)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1096
  apply (drule_tac x = i in bspec, assumption)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1097
  apply (subgoal_tac "(UN y:(B i). {(i, y)}) <= (%y. (i, y)) ` (B i)") 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1098
  apply (rule finite_surj)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1099
  apply auto
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1100
  apply (rule setprod_cong, rule refl)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1101
  apply (subst setprod_UN_disjoint)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1102
  apply (erule bspec, assumption)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1103
  apply auto
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1104
  done
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1105
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1106
lemma setprod_cartesian_product: "finite A ==> finite B ==> 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1107
    (\<Prod> x:A. (\<Prod> y: B. f x y)) = 
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1108
    (\<Prod> z:(A <*> B). f (fst z) (snd z))"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1109
  by (erule setprod_Sigma, auto)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1110
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1111
lemma setprod_timesf: "setprod (%x. f x * g x) A =
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1112
    (setprod f A * setprod g A)"
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1113
  apply (case_tac "finite A")
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1114
   prefer 2 apply (simp add: setprod_def times_ac1)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1115
  apply (erule finite_induct, auto)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1116
  apply (simp add: times_ac1)
ea2707645af8 new material from Avigad
paulson
parents: 14449
diff changeset
  1117
  done
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1118
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1119
subsubsection {* Properties in more restricted classes of structures *}
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1120
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1121
lemma setprod_eq_1_iff [simp]:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1122
    "finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1123
  by (induct set: Finites) auto
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1124
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1125
lemma setprod_constant: "finite A ==> (\<Prod>x: A. (y::'a::ringpower)) =
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1126
    y^(card A)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1127
  apply (erule finite_induct)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1128
  apply (auto simp add: power_Suc)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1129
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1130
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1131
lemma setprod_zero: "finite A ==> EX x: A. f x = (0::'a::semiring) ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1132
    setprod f A = 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1133
  apply (induct set: Finites, force, clarsimp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1134
  apply (erule disjE, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1135
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1136
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1137
lemma setprod_nonneg [rule_format]: "(ALL x: A. (0::'a::ordered_ring) \<le> f x)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1138
     --> 0 \<le> setprod f A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1139
  apply (case_tac "finite A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1140
  apply (induct set: Finites, force, clarsimp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1141
  apply (subgoal_tac "0 * 0 \<le> f x * setprod f F", force)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1142
  apply (rule mult_mono, assumption+)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1143
  apply (auto simp add: setprod_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1144
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1145
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1146
lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::ordered_ring) < f x)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1147
     --> 0 < setprod f A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1148
  apply (case_tac "finite A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1149
  apply (induct set: Finites, force, clarsimp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1150
  apply (subgoal_tac "0 * 0 < f x * setprod f F", force)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1151
  apply (rule mult_strict_mono, assumption+)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1152
  apply (auto simp add: setprod_def)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1153
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1154
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1155
lemma setprod_nonzero [rule_format]:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1156
    "(ALL x y. (x::'a::semiring) * y = 0 --> x = 0 | y = 0) ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1157
      finite A ==> (ALL x: A. f x \<noteq> (0::'a)) --> setprod f A \<noteq> 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1158
  apply (erule finite_induct, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1159
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1160
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1161
lemma setprod_zero_eq:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1162
    "(ALL x y. (x::'a::semiring) * y = 0 --> x = 0 | y = 0) ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1163
     finite A ==> (setprod f A = (0::'a)) = (EX x: A. f x = 0)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1164
  apply (insert setprod_zero [of A f] setprod_nonzero [of A f], blast)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1165
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1166
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1167
lemma setprod_nonzero_field:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1168
    "finite A ==> (ALL x: A. f x \<noteq> (0::'a::field)) ==> setprod f A \<noteq> 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1169
  apply (rule setprod_nonzero, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1170
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1171
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1172
lemma setprod_zero_eq_field:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1173
    "finite A ==> (setprod f A = (0::'a::field)) = (EX x: A. f x = 0)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1174
  apply (rule setprod_zero_eq, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1175
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1176
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1177
lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1178
    (setprod f (A Un B) :: 'a ::{field})
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1179
      = setprod f A * setprod f B / setprod f (A Int B)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1180
  apply (subst setprod_Un_Int [symmetric], auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1181
  apply (subgoal_tac "finite (A Int B)")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1182
  apply (frule setprod_nonzero_field [of "A Int B" f], assumption)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1183
  apply (subst times_divide_eq_right [THEN sym], auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1184
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1185
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1186
lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1187
    (setprod f (A - {a}) :: 'a :: {field}) =
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1188
      (if a:A then setprod f A / f a else setprod f A)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1189
  apply (erule finite_induct)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1190
   apply (auto simp add: insert_Diff_if)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1191
  apply (subgoal_tac "f a * setprod f F / f a = setprod f F * f a / f a")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1192
  apply (erule ssubst)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1193
  apply (subst times_divide_eq_right [THEN sym])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1194
  apply (auto simp add: mult_ac)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1195
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1196
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1197
lemma setprod_inversef: "finite A ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1198
    ALL x: A. f x \<noteq> (0::'a::{field,division_by_zero}) ==>
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1199
      setprod (inverse \<circ> f) A = inverse (setprod f A)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1200
  apply (erule finite_induct)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1201
  apply (simp, simp)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1202
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1203
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1204
lemma setprod_dividef:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1205
     "[|finite A;
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1206
        \<forall>x \<in> A. g x \<noteq> (0::'a::{field,division_by_zero})|]
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1207
      ==> setprod (%x. f x / g x) A = setprod f A / setprod g A"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1208
  apply (subgoal_tac
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1209
         "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A")
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1210
  apply (erule ssubst)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1211
  apply (subst divide_inverse)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1212
  apply (subst setprod_timesf)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1213
  apply (subst setprod_inversef, assumption+, rule refl)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1214
  apply (rule setprod_cong, rule refl)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1215
  apply (subst divide_inverse, auto)
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1216
  done
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1217
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1218
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1219
subsection{* Min and Max of finite linearly ordered sets *}
13490
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1220
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1221
text{* Seemed easier to define directly than via fold. *}
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1222
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1223
lemma ex_Max: fixes S :: "('a::linorder)set"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1224
  assumes fin: "finite S" shows "S \<noteq> {} ==> \<exists>m\<in>S. \<forall>s \<in> S. s \<le> m"
13490
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1225
using fin
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1226
proof (induct)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1227
  case empty thus ?case by simp
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1228
next
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1229
  case (insert S x)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1230
  show ?case
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1231
  proof (cases)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1232
    assume "S = {}" thus ?thesis by simp
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1233
  next
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1234
    assume nonempty: "S \<noteq> {}"
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1235
    then obtain m where m: "m\<in>S" "\<forall>s\<in>S. s \<le> m" using insert by blast
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1236
    show ?thesis
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1237
    proof (cases)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1238
      assume "x \<le> m" thus ?thesis using m by blast
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1239
    next
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1240
      assume "~ x \<le> m" thus ?thesis using m
13490
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1241
	by(simp add:linorder_not_le order_less_le)(blast intro: order_trans)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1242
    qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1243
  qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1244
qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1245
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1246
lemma ex_Min: fixes S :: "('a::linorder)set"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1247
  assumes fin: "finite S" shows "S \<noteq> {} ==> \<exists>m\<in>S. \<forall>s \<in> S. m \<le> s"
13490
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1248
using fin
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1249
proof (induct)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1250
  case empty thus ?case by simp
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1251
next
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1252
  case (insert S x)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1253
  show ?case
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1254
  proof (cases)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1255
    assume "S = {}" thus ?thesis by simp
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1256
  next
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1257
    assume nonempty: "S \<noteq> {}"
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1258
    then obtain m where m: "m\<in>S" "\<forall>s\<in>S. m \<le> s" using insert by blast
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1259
    show ?thesis
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1260
    proof (cases)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1261
      assume "m \<le> x" thus ?thesis using m by blast
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1262
    next
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1263
      assume "~ m \<le> x" thus ?thesis using m
13490
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1264
	by(simp add:linorder_not_le order_less_le)(blast intro: order_trans)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1265
    qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1266
  qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1267
qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1268
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1269
constdefs
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1270
 Min :: "('a::linorder)set => 'a"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1271
"Min S  ==  THE m. m \<in> S \<and> (\<forall>s \<in> S. m \<le> s)"
13490
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1272
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1273
 Max :: "('a::linorder)set => 'a"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1274
"Max S  ==  THE m. m \<in> S \<and> (\<forall>s \<in> S. s \<le> m)"
13490
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1275
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1276
lemma Min[simp]: assumes a: "finite S" "S \<noteq> {}"
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1277
                 shows "Min S \<in> S \<and> (\<forall>s \<in> S. Min S \<le> s)" (is "?P(Min S)")
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1278
proof (unfold Min_def, rule theI')
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1279
  show "\<exists>!m. ?P m"
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1280
  proof (rule ex_ex1I)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1281
    show "\<exists>m. ?P m" using ex_Min[OF a] by blast
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1282
  next
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1283
    fix m1 m2 assume "?P m1" "?P m2"
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1284
    thus "m1 = m2" by (blast dest:order_antisym)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1285
  qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1286
qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1287
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1288
lemma Max[simp]: assumes a: "finite S" "S \<noteq> {}"
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1289
                 shows "Max S \<in> S \<and> (\<forall>s \<in> S. s \<le> Max S)" (is "?P(Max S)")
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1290
proof (unfold Max_def, rule theI')
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1291
  show "\<exists>!m. ?P m"
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1292
  proof (rule ex_ex1I)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1293
    show "\<exists>m. ?P m" using ex_Max[OF a] by blast
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1294
  next
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1295
    fix m1 m2 assume "?P m1" "?P m2"
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1296
    thus "m1 = m2" by (blast dest:order_antisym)
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1297
  qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1298
qed
44bdc150211b Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents: 13421
diff changeset
  1299
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1300
subsection {* Theorems about @{text "choose"} *}
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1301
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1302
text {*
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1303
  \medskip Basic theorem about @{text "choose"}.  By Florian
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1304
  Kammüller, tidied by LCP.
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1305
*}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1306
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1307
lemma card_s_0_eq_empty:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1308
    "finite A ==> card {B. B \<subseteq> A & card B = 0} = 1"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1309
  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
  1310
  apply (simp cong add: rev_conj_cong)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1311
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1312
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1313
lemma choose_deconstruct: "finite M ==> x \<notin> M
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1314
  ==> {s. s <= insert x M & card(s) = Suc k}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1315
       = {s. s <= M & card(s) = Suc k} Un
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1316
         {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
  1317
  apply safe
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1318
   apply (auto intro: finite_subset [THEN card_insert_disjoint])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1319
  apply (drule_tac x = "xa - {x}" in spec)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1320
  apply (subgoal_tac "x \<notin> xa", auto)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1321
  apply (erule rev_mp, subst card_Diff_singleton)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1322
  apply (auto intro: finite_subset)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1323
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1324
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1325
lemma card_inj_on_le:
13595
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1326
    "[|inj_on f A; f ` A \<subseteq> B; finite A; finite B |] ==> card A <= card B"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1327
  by (auto intro: card_mono simp add: card_image [symmetric])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1328
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1329
lemma card_bij_eq:
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14331
diff changeset
  1330
    "[|inj_on f A; f ` A \<subseteq> B; inj_on g B; g ` B \<subseteq> A;
13595
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1331
       finite A; finite B |] ==> card A = card B"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1332
  by (auto intro: le_anti_sym card_inj_on_le)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1333
13595
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1334
text{*There are as many subsets of @{term A} having cardinality @{term k}
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1335
 as there are sets obtained from the former by inserting a fixed element
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1336
 @{term x} into each.*}
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1337
lemma constr_bij:
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1338
   "[|finite A; x \<notin> A|] ==>
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1339
    card {B. EX C. C <= A & card(C) = k & B = insert x C} =
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1340
    card {B. B <= A & card(B) = k}"
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1341
  apply (rule_tac f = "%s. s - {x}" and g = "insert x" in card_bij_eq)
13595
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1342
       apply (auto elim!: equalityE simp add: inj_on_def)
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1343
    apply (subst Diff_insert0, auto)
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1344
   txt {* finiteness of the two sets *}
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1345
   apply (rule_tac [2] B = "Pow (A)" in finite_subset)
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1346
   apply (rule_tac B = "Pow (insert x A)" in finite_subset)
7e6cdcd113a2 Proof tidying
paulson
parents: 13571
diff changeset
  1347
   apply fast+
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1348
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1349
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1350
text {*
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1351
  Main theorem: combinatorial statement about number of subsets of a set.
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1352
*}
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1353
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1354
lemma n_sub_lemma:
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1355
  "!!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
  1356
  apply (induct k)
14208
144f45277d5a misc tidying
paulson
parents: 13825
diff changeset
  1357
   apply (simp add: card_s_0_eq_empty, atomize)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1358
  apply (rotate_tac -1, erule finite_induct)
13421
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
  1359
   apply (simp_all (no_asm_simp) cong add: conj_cong
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
  1360
     add: card_s_0_eq_empty choose_deconstruct)
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1361
  apply (subst card_Un_disjoint)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1362
     prefer 4 apply (force simp add: constr_bij)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1363
    prefer 3 apply force
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1364
   prefer 2 apply (blast intro: finite_Pow_iff [THEN iffD2]
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1365
     finite_subset [of _ "Pow (insert x F)", standard])
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1366
  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
  1367
  done
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1368
13421
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
  1369
theorem n_subsets:
8fcdf4a26468 simplified locale predicates;
wenzelm
parents: 13400
diff changeset
  1370
    "finite A ==> card {B. B <= A & card B = k} = (card A choose k)"
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1371
  by (simp add: n_sub_lemma)
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1372
14443
75910c7557c5 generic theorems about exponentials; general tidying up
paulson
parents: 14430
diff changeset
  1373
12396
2298d5b8e530 renamed theory Finite to Finite_Set and converted;
wenzelm
parents:
diff changeset
  1374
end