src/HOL/Binomial.thy
author wenzelm
Wed, 11 Dec 2024 11:18:52 +0100
changeset 81579 cf4bebd770b5
parent 80932 261cd8722677
permissions -rw-r--r--
proper bundle binomial_syntax; NB: precedence of "choose" changes silently from 65 to 64 in 200107cdd3ac, but old 65 was still seen in the wild;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
     1
(*  Title:      HOL/Binomial.thy
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
     2
    Author:     Jacques D. Fleuriot
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
     3
    Author:     Lawrence C Paulson
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
     4
    Author:     Jeremy Avigad
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
     5
    Author:     Chaitanya Mangla
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
     6
    Author:     Manuel Eberl
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     7
*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     8
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
     9
section \<open>Binomial Coefficients, Binomial Theorem, Inclusion-exclusion Principle\<close>
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
    10
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
    11
theory Binomial
65813
bdd17b18e103 relaxed theory dependencies
haftmann
parents: 65812
diff changeset
    12
  imports Presburger Factorial
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
    13
begin
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
    14
63373
487d764fca4a tuned sections
haftmann
parents: 63372
diff changeset
    15
subsection \<open>Binomial coefficients\<close>
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    16
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
    17
text \<open>This development is based on the work of Andy Gordon and Florian Kammueller.\<close>
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
    18
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    19
text \<open>Combinatorial definition\<close>
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    20
81579
cf4bebd770b5 proper bundle binomial_syntax;
wenzelm
parents: 80932
diff changeset
    21
definition binomial :: "nat \<Rightarrow> nat \<Rightarrow> nat"
cf4bebd770b5 proper bundle binomial_syntax;
wenzelm
parents: 80932
diff changeset
    22
  where "binomial n k = card {K\<in>Pow {0..<n}. card K = k}"
cf4bebd770b5 proper bundle binomial_syntax;
wenzelm
parents: 80932
diff changeset
    23
cf4bebd770b5 proper bundle binomial_syntax;
wenzelm
parents: 80932
diff changeset
    24
open_bundle binomial_syntax
cf4bebd770b5 proper bundle binomial_syntax;
wenzelm
parents: 80932
diff changeset
    25
begin
cf4bebd770b5 proper bundle binomial_syntax;
wenzelm
parents: 80932
diff changeset
    26
notation binomial  (infix \<open>choose\<close> 64)
cf4bebd770b5 proper bundle binomial_syntax;
wenzelm
parents: 80932
diff changeset
    27
end
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    28
80175
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 79586
diff changeset
    29
lemma binomial_right_mono:
79586
9cde97e471df Two new theorems
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
    30
  assumes "m \<le> n" shows "m choose k \<le> n choose k"
9cde97e471df Two new theorems
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
    31
proof -
9cde97e471df Two new theorems
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
    32
  have "{K. K \<subseteq> {0..<m} \<and> card K = k} \<subseteq> {K. K \<subseteq> {0..<n} \<and> card K = k}"
9cde97e471df Two new theorems
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
    33
    using assms by auto
9cde97e471df Two new theorems
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
    34
  then show ?thesis
9cde97e471df Two new theorems
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
    35
    by (simp add: binomial_def card_mono)
9cde97e471df Two new theorems
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
    36
qed
9cde97e471df Two new theorems
paulson <lp15@cam.ac.uk>
parents: 79566
diff changeset
    37
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    38
theorem n_subsets:
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    39
  assumes "finite A"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    40
  shows "card {B. B \<subseteq> A \<and> card B = k} = card A choose k"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    41
proof -
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    42
  from assms obtain f where bij: "bij_betw f {0..<card A} A"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    43
    by (blast dest: ex_bij_betw_nat_finite)
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    44
  then have [simp]: "card (f ` C) = card C" if "C \<subseteq> {0..<card A}" for C
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    45
    by (meson bij_betw_imp_inj_on bij_betw_subset card_image that)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    46
  from bij have "bij_betw (image f) (Pow {0..<card A}) (Pow A)"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    47
    by (rule bij_betw_Pow)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    48
  then have "inj_on (image f) (Pow {0..<card A})"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    49
    by (rule bij_betw_imp_inj_on)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    50
  moreover have "{K. K \<subseteq> {0..<card A} \<and> card K = k} \<subseteq> Pow {0..<card A}"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    51
    by auto
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    52
  ultimately have "inj_on (image f) {K. K \<subseteq> {0..<card A} \<and> card K = k}"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    53
    by (rule inj_on_subset)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    54
  then have "card {K. K \<subseteq> {0..<card A} \<and> card K = k} =
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
    55
      card (image f ` {K. K \<subseteq> {0..<card A} \<and> card K = k})" (is "_ = card ?C")
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    56
    by (simp add: card_image)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    57
  also have "?C = {K. K \<subseteq> f ` {0..<card A} \<and> card K = k}"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    58
    by (auto elim!: subset_imageE)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    59
  also have "f ` {0..<card A} = A"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    60
    by (meson bij bij_betw_def)
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    61
  finally show ?thesis
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    62
    by (simp add: binomial_def)
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    63
qed
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
    64
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    65
text \<open>Recursive characterization\<close>
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    66
68785
862b1288020f tuned code setup
haftmann
parents: 68784
diff changeset
    67
lemma binomial_n_0 [simp]: "n choose 0 = 1"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    68
proof -
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    69
  have "{K \<in> Pow {0..<n}. card K = 0} = {{}}"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    70
    by (auto dest: finite_subset)
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    71
  then show ?thesis
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    72
    by (simp add: binomial_def)
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    73
qed
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    74
68785
862b1288020f tuned code setup
haftmann
parents: 68784
diff changeset
    75
lemma binomial_0_Suc [simp]: "0 choose Suc k = 0"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    76
  by (simp add: binomial_def)
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
    77
68785
862b1288020f tuned code setup
haftmann
parents: 68784
diff changeset
    78
lemma binomial_Suc_Suc [simp]: "Suc n choose Suc k = (n choose k) + (n choose Suc k)"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    79
proof -
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    80
  let ?P = "\<lambda>n k. {K. K \<subseteq> {0..<n} \<and> card K = k}"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    81
  let ?Q = "?P (Suc n) (Suc k)"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    82
  have inj: "inj_on (insert n) (?P n k)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
    83
    by rule (auto; metis atLeastLessThan_iff insert_iff less_irrefl subsetCE)
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    84
  have disjoint: "insert n ` ?P n k \<inter> ?P n (Suc k) = {}"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    85
    by auto
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    86
  have "?Q = {K\<in>?Q. n \<in> K} \<union> {K\<in>?Q. n \<notin> K}"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    87
    by auto
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    88
  also have "{K\<in>?Q. n \<in> K} = insert n ` ?P n k" (is "?A = ?B")
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    89
  proof (rule set_eqI)
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    90
    fix K
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
    91
    have K_finite: "finite K" if "K \<subseteq> insert n {0..<n}"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    92
      using that by (rule finite_subset) simp_all
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    93
    have Suc_card_K: "Suc (card K - Suc 0) = card K" if "n \<in> K"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    94
      and "finite K"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    95
    proof -
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    96
      from \<open>n \<in> K\<close> obtain L where "K = insert n L" and "n \<notin> L"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    97
        by (blast elim: Set.set_insert)
72302
d7d90ed4c74e fixed some remarkably ugly proofs
paulson <lp15@cam.ac.uk>
parents: 71720
diff changeset
    98
      with that show ?thesis by (simp add: card.insert_remove)
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
    99
    qed
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   100
    show "K \<in> ?A \<longleftrightarrow> K \<in> ?B"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   101
      by (subst in_image_insert_iff)
72302
d7d90ed4c74e fixed some remarkably ugly proofs
paulson <lp15@cam.ac.uk>
parents: 71720
diff changeset
   102
        (auto simp add: card.insert_remove subset_eq_atLeast0_lessThan_finite
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   103
          Diff_subset_conv K_finite Suc_card_K)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   104
  qed
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   105
  also have "{K\<in>?Q. n \<notin> K} = ?P n (Suc k)"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   106
    by (auto simp add: atLeast0_lessThan_Suc)
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   107
  finally show ?thesis using inj disjoint
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   108
    by (simp add: binomial_def card_Un_disjoint card_image)
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   109
qed
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   110
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   111
lemma binomial_eq_0: "n < k \<Longrightarrow> n choose k = 0"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   112
  by (auto simp add: binomial_def dest: subset_eq_atLeast0_lessThan_card)
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   113
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   114
lemma zero_less_binomial: "k \<le> n \<Longrightarrow> n choose k > 0"
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   115
  by (induct n k rule: diff_induct) simp_all
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   116
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   117
lemma binomial_eq_0_iff [simp]: "n choose k = 0 \<longleftrightarrow> n < k"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   118
  by (metis binomial_eq_0 less_numeral_extra(3) not_less zero_less_binomial)
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   119
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   120
lemma zero_less_binomial_iff [simp]: "n choose k > 0 \<longleftrightarrow> k \<le> n"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   121
  by (metis binomial_eq_0_iff not_less0 not_less zero_less_binomial)
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   122
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   123
lemma binomial_n_n [simp]: "n choose n = 1"
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   124
  by (induct n) (simp_all add: binomial_eq_0)
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   125
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   126
lemma binomial_Suc_n [simp]: "Suc n choose n = Suc n"
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   127
  by (induct n) simp_all
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   128
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   129
lemma binomial_1 [simp]: "n choose Suc 0 = n"
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   130
  by (induct n) simp_all
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   131
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
   132
lemma choose_one: "n choose 1 = n" for n :: nat
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
   133
  by simp
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
   134
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   135
lemma choose_reduce_nat:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   136
  "0 < n \<Longrightarrow> 0 < k \<Longrightarrow>
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   137
    n choose k = ((n - 1) choose (k - 1)) + ((n - 1) choose k)"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   138
  using binomial_Suc_Suc [of "n - 1" "k - 1"] by simp
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   139
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   140
lemma Suc_times_binomial_eq: "Suc n * (n choose k) = (Suc n choose Suc k) * Suc k"
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   141
proof (induction n arbitrary: k)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   142
  case 0
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   143
  then show ?case
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   144
    by auto
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   145
next
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   146
  case (Suc n)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   147
  show ?case 
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   148
  proof (cases k)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   149
    case (Suc k')
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   150
    then show ?thesis
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   151
      using Suc.IH
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   152
      by (auto simp add: add_mult_distrib add_mult_distrib2 le_Suc_eq binomial_eq_0)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   153
  qed auto
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   154
qed
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   155
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   156
lemma binomial_le_pow2: "n choose k \<le> 2^n"
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   157
proof (induction n arbitrary: k)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   158
  case 0
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   159
  then show ?case
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   160
    using le_less less_le_trans by fastforce
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   161
next
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   162
  case (Suc n)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   163
  show ?case
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   164
  proof (cases k)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   165
    case (Suc k')
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   166
    then show ?thesis
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   167
      using Suc.IH by (simp add: add_le_mono mult_2)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   168
  qed auto
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   169
qed
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   170
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   171
text \<open>The absorption property.\<close>
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   172
lemma Suc_times_binomial: "Suc k * (Suc n choose Suc k) = Suc n * (n choose k)"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   173
  using Suc_times_binomial_eq by auto
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   174
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   175
text \<open>This is the well-known version of absorption, but it's harder to use
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   176
  because of the need to reason about division.\<close>
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   177
lemma binomial_Suc_Suc_eq_times: "(Suc n choose Suc k) = (Suc n * (n choose k)) div Suc k"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   178
  by (simp add: Suc_times_binomial_eq del: mult_Suc mult_Suc_right)
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   179
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   180
text \<open>Another version of absorption, with \<open>-1\<close> instead of \<open>Suc\<close>.\<close>
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   181
lemma times_binomial_minus1_eq: "0 < k \<Longrightarrow> k * (n choose k) = n * ((n - 1) choose (k - 1))"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   182
  using Suc_times_binomial_eq [where n = "n - 1" and k = "k - 1"]
63648
f9f3006a5579 "split add" -> "split"
nipkow
parents: 63526
diff changeset
   183
  by (auto split: nat_diff_split)
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   184
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   185
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60604
diff changeset
   186
subsection \<open>The binomial theorem (courtesy of Tobias Nipkow):\<close>
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   187
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   188
text \<open>Avigad's version, generalized to any commutative ring\<close>
79544
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 78667
diff changeset
   189
theorem (in comm_semiring_1) binomial_ring:
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 78667
diff changeset
   190
  "(a + b :: 'a)^n = (\<Sum>k\<le>n. (of_nat (n choose k)) * a^k * b^(n-k))"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   191
proof (induct n)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   192
  case 0
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   193
  then show ?case by simp
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   194
next
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   195
  case (Suc n)
79544
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 78667
diff changeset
   196
  have decomp: "{0..n+1} = {0} \<union> {n + 1} \<union> {1..n}" and decomp2: "{0..n} = {0} \<union> {1..n}"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   197
    by auto
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   198
  have "(a + b)^(n+1) = (a + b) * (\<Sum>k\<le>n. of_nat (n choose k) * a^k * b^(n - k))"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   199
    using Suc.hyps by simp
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   200
  also have "\<dots> = a * (\<Sum>k\<le>n. of_nat (n choose k) * a^k * b^(n-k)) +
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   201
      b * (\<Sum>k\<le>n. of_nat (n choose k) * a^k * b^(n-k))"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   202
    by (rule distrib_right)
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   203
  also have "\<dots> = (\<Sum>k\<le>n. of_nat (n choose k) * a^(k+1) * b^(n-k)) +
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   204
      (\<Sum>k\<le>n. of_nat (n choose k) * a^k * b^(n - k + 1))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   205
    by (auto simp add: sum_distrib_left ac_simps)
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   206
  also have "\<dots> = (\<Sum>k\<le>n. of_nat (n choose k) * a^k * b^(n + 1 - k)) +
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   207
      (\<Sum>k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n + 1 - k))"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   208
    by (simp add: atMost_atLeast0 sum.shift_bounds_cl_Suc_ivl Suc_diff_le field_simps del: sum.cl_ivl_Suc)
71351
b3a93a91803b generalized thm (as suggested by Christian Weinz)
nipkow
parents: 70113
diff changeset
   209
  also have "\<dots> = b^(n + 1) +
b3a93a91803b generalized thm (as suggested by Christian Weinz)
nipkow
parents: 70113
diff changeset
   210
      (\<Sum>k=1..n. of_nat (n choose k) * a^k * b^(n + 1 - k)) + (a^(n + 1) +
b3a93a91803b generalized thm (as suggested by Christian Weinz)
nipkow
parents: 70113
diff changeset
   211
      (\<Sum>k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n + 1 - k)))"
b3a93a91803b generalized thm (as suggested by Christian Weinz)
nipkow
parents: 70113
diff changeset
   212
      using sum.nat_ivl_Suc' [of 1 n "\<lambda>k. of_nat (n choose (k-1)) * a ^ k * b ^ (n + 1 - k)"]
b3a93a91803b generalized thm (as suggested by Christian Weinz)
nipkow
parents: 70113
diff changeset
   213
    by (simp add: sum.atLeast_Suc_atMost atMost_atLeast0)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   214
  also have "\<dots> = a^(n + 1) + b^(n + 1) +
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   215
      (\<Sum>k=1..n. of_nat (n + 1 choose k) * a^k * b^(n + 1 - k))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   216
    by (auto simp add: field_simps sum.distrib [symmetric] choose_reduce_nat)
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   217
  also have "\<dots> = (\<Sum>k\<le>n+1. of_nat (n + 1 choose k) * a^k * b^(n + 1 - k))"
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   218
    using decomp by (simp add: atMost_atLeast0 field_simps)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   219
  finally show ?case
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   220
    by simp
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   221
qed
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   222
79544
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 78667
diff changeset
   223
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   224
text \<open>Original version for the naturals.\<close>
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   225
corollary binomial: "(a + b :: nat)^n = (\<Sum>k\<le>n. (of_nat (n choose k)) * a^k * b^(n - k))"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   226
  using binomial_ring [of "int a" "int b" n]
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   227
  by (simp only: of_nat_add [symmetric] of_nat_mult [symmetric] of_nat_power [symmetric]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   228
      of_nat_sum [symmetric] of_nat_eq_iff of_nat_id)
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   229
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   230
lemma binomial_fact_lemma: "k \<le> n \<Longrightarrow> fact k * fact (n - k) * (n choose k) = fact n"
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   231
proof (induct n arbitrary: k rule: nat_less_induct)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   232
  fix n k
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   233
  assume H: "\<forall>m<n. \<forall>x\<le>m. fact x * fact (m - x) * (m choose x) = fact m"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   234
  assume kn: "k \<le> n"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   235
  let ?ths = "fact k * fact (n - k) * (n choose k) = fact n"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   236
  consider "n = 0 \<or> k = 0 \<or> n = k" | m h where "n = Suc m" "k = Suc h" "h < m"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   237
    using kn by atomize_elim presburger
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   238
  then show "fact k * fact (n - k) * (n choose k) = fact n"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   239
  proof cases
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   240
    case 1
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   241
    with kn show ?thesis by auto
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   242
  next
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   243
    case 2
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   244
    note n = \<open>n = Suc m\<close>
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   245
    note k = \<open>k = Suc h\<close>
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   246
    note hm = \<open>h < m\<close>
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   247
    have mn: "m < n"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   248
      using n by arith
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   249
    have hm': "h \<le> m"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   250
      using hm by arith
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   251
    have km: "k \<le> m"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   252
      using hm k n kn by arith
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   253
    have "m - h = Suc (m - Suc h)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   254
      using  k km hm by arith
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   255
    with km k have "fact (m - h) = (m - h) * fact (m - k)"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   256
      by simp
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   257
    with n k have "fact k * fact (n - k) * (n choose k) =
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   258
        k * (fact h * fact (m - h) * (m choose h)) +
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   259
        (m - h) * (fact k * fact (m - k) * (m choose k))"
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   260
      by (simp add: field_simps)
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   261
    also have "\<dots> = (k + (m - h)) * fact m"
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   262
      using H[rule_format, OF mn hm'] H[rule_format, OF mn km]
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   263
      by (simp add: field_simps)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   264
    finally show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   265
      using k n km by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   266
  qed
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   267
qed
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   268
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   269
lemma binomial_fact':
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   270
  assumes "k \<le> n"
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   271
  shows "n choose k = fact n div (fact k * fact (n - k))"
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   272
  using binomial_fact_lemma [OF assms]
64240
eabf80376aab more standardized names
haftmann
parents: 63918
diff changeset
   273
  by (metis fact_nonzero mult_eq_0_iff nonzero_mult_div_cancel_left)
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   274
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   275
lemma binomial_fact:
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   276
  assumes kn: "k \<le> n"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   277
  shows "(of_nat (n choose k) :: 'a::field_char_0) = fact n / (fact k * fact (n - k))"
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   278
  using binomial_fact_lemma[OF kn]
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   279
  by (metis (mono_tags, lifting) fact_nonzero mult_eq_0_iff nonzero_mult_div_cancel_left of_nat_fact of_nat_mult)
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58889
diff changeset
   280
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   281
lemma fact_binomial:
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   282
  assumes "k \<le> n"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   283
  shows "fact k * of_nat (n choose k) = (fact n / fact (n - k) :: 'a::field_char_0)"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   284
  unfolding binomial_fact [OF assms] by (simp add: field_simps)
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   285
79566
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79544
diff changeset
   286
lemma binomial_fact_pow: "(n choose s) * fact s \<le> n^s"
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79544
diff changeset
   287
proof (cases "s \<le> n")
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79544
diff changeset
   288
  case True
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79544
diff changeset
   289
  then show ?thesis
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79544
diff changeset
   290
    by (smt (verit) binomial_fact_lemma mult.assoc mult.commute fact_div_fact_le_pow fact_nonzero nonzero_mult_div_cancel_right) 
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79544
diff changeset
   291
qed (simp add: binomial_eq_0)
f783490c6c99 A small number of new lemmas
paulson <lp15@cam.ac.uk>
parents: 79544
diff changeset
   292
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   293
lemma choose_two: "n choose 2 = n * (n - 1) div 2"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   294
proof (cases "n \<ge> 2")
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   295
  case False
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   296
  then have "n = 0 \<or> n = 1"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   297
    by auto
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   298
  then show ?thesis by auto
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   299
next
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   300
  case True
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   301
  define m where "m = n - 2"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   302
  with True have "n = m + 2"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   303
    by simp
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   304
  then have "fact n = n * (n - 1) * fact (n - 2)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   305
    by (simp add: fact_prod_Suc atLeast0_lessThan_Suc algebra_simps)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   306
  with True show ?thesis
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   307
    by (simp add: binomial_fact')
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   308
qed
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   309
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   310
lemma choose_row_sum: "(\<Sum>k\<le>n. n choose k) = 2^n"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   311
  using binomial [of 1 "1" n] by (simp add: numeral_2_eq_2)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   312
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   313
lemma sum_choose_lower: "(\<Sum>k\<le>n. (r+k) choose k) = Suc (r+n) choose n"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   314
  by (induct n) auto
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   315
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   316
lemma sum_choose_upper: "(\<Sum>k\<le>n. k choose m) = Suc n choose Suc m"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   317
  by (induct n) auto
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   318
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   319
lemma choose_alternating_sum:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   320
  "n > 0 \<Longrightarrow> (\<Sum>i\<le>n. (-1)^i * of_nat (n choose i)) = (0 :: 'a::comm_ring_1)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   321
  using binomial_ring[of "-1 :: 'a" 1 n]
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   322
  by (simp add: atLeast0AtMost mult_of_nat_commute zero_power)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   323
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   324
lemma choose_even_sum:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   325
  assumes "n > 0"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   326
  shows "2 * (\<Sum>i\<le>n. if even i then of_nat (n choose i) else 0) = (2 ^ n :: 'a::comm_ring_1)"
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   327
proof -
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   328
  have "2 ^ n = (\<Sum>i\<le>n. of_nat (n choose i)) + (\<Sum>i\<le>n. (-1) ^ i * of_nat (n choose i) :: 'a)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   329
    using choose_row_sum[of n]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   330
    by (simp add: choose_alternating_sum assms atLeast0AtMost of_nat_sum[symmetric])
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   331
  also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) + (-1) ^ i * of_nat (n choose i))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   332
    by (simp add: sum.distrib)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   333
  also have "\<dots> = 2 * (\<Sum>i\<le>n. if even i then of_nat (n choose i) else 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   334
    by (subst sum_distrib_left, intro sum.cong) simp_all
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   335
  finally show ?thesis ..
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   336
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   337
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   338
lemma choose_odd_sum:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   339
  assumes "n > 0"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   340
  shows "2 * (\<Sum>i\<le>n. if odd i then of_nat (n choose i) else 0) = (2 ^ n :: 'a::comm_ring_1)"
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   341
proof -
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   342
  have "2 ^ n = (\<Sum>i\<le>n. of_nat (n choose i)) - (\<Sum>i\<le>n. (-1) ^ i * of_nat (n choose i) :: 'a)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   343
    using choose_row_sum[of n]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   344
    by (simp add: choose_alternating_sum assms atLeast0AtMost of_nat_sum[symmetric])
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   345
  also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) - (-1) ^ i * of_nat (n choose i))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   346
    by (simp add: sum_subtractf)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   347
  also have "\<dots> = 2 * (\<Sum>i\<le>n. if odd i then of_nat (n choose i) else 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   348
    by (subst sum_distrib_left, intro sum.cong) simp_all
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   349
  finally show ?thesis ..
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   350
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   351
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60604
diff changeset
   352
text\<open>NW diagonal sum property\<close>
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   353
lemma sum_choose_diagonal:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   354
  assumes "m \<le> n"
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   355
  shows "(\<Sum>k\<le>m. (n - k) choose (m - k)) = Suc n choose m"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   356
proof -
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   357
  have "(\<Sum>k\<le>m. (n-k) choose (m - k)) = (\<Sum>k\<le>m. (n - m + k) choose k)"
67411
3f4b0c84630f restored naming of lemmas after corresponding constants
haftmann
parents: 67399
diff changeset
   358
    using sum.atLeastAtMost_rev [of "\<lambda>k. (n - k) choose (m - k)" 0 m] assms
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   359
    by (simp add: atMost_atLeast0)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   360
  also have "\<dots> = Suc (n - m + m) choose m"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   361
    by (rule sum_choose_lower)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   362
  also have "\<dots> = Suc n choose m"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   363
    using assms by simp
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   364
  finally show ?thesis .
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   365
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   366
63373
487d764fca4a tuned sections
haftmann
parents: 63372
diff changeset
   367
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   368
subsection \<open>Generalized binomial coefficients\<close>
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   369
80932
261cd8722677 standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents: 80176
diff changeset
   370
definition gbinomial :: "'a::{semidom_divide,semiring_char_0} \<Rightarrow> nat \<Rightarrow> 'a"  (infix \<open>gchoose\<close> 64)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   371
  where gbinomial_prod_rev: "a gchoose k = prod (\<lambda>i. a - of_nat i) {0..<k} div fact k"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   372
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   373
lemma gbinomial_0 [simp]:
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   374
  "a gchoose 0 = 1"
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   375
  "0 gchoose (Suc k) = 0"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   376
  by (simp_all add: gbinomial_prod_rev prod.atLeast0_lessThan_Suc_shift del: prod.op_ivl_Suc)
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63366
diff changeset
   377
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   378
lemma gbinomial_Suc: "a gchoose (Suc k) = prod (\<lambda>i. a - of_nat i) {0..k} div fact (Suc k)"
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   379
  by (simp add: gbinomial_prod_rev atLeastLessThanSuc_atLeastAtMost)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   380
68786
134be95e5787 slightly generalized theorems
haftmann
parents: 68785
diff changeset
   381
lemma gbinomial_1 [simp]: "a gchoose 1 = a"
134be95e5787 slightly generalized theorems
haftmann
parents: 68785
diff changeset
   382
  by (simp add: gbinomial_prod_rev lessThan_Suc)
134be95e5787 slightly generalized theorems
haftmann
parents: 68785
diff changeset
   383
134be95e5787 slightly generalized theorems
haftmann
parents: 68785
diff changeset
   384
lemma gbinomial_Suc0 [simp]: "a gchoose Suc 0 = a"
134be95e5787 slightly generalized theorems
haftmann
parents: 68785
diff changeset
   385
  by (simp add: gbinomial_prod_rev lessThan_Suc)
134be95e5787 slightly generalized theorems
haftmann
parents: 68785
diff changeset
   386
80176
7fefa7839ac6 syntax of gchoose now the same as choose
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
   387
lemma gbinomial_0_left: "0 gchoose k = (if k = 0 then 1 else 0)"
7fefa7839ac6 syntax of gchoose now the same as choose
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
   388
  by (cases k) simp_all
7fefa7839ac6 syntax of gchoose now the same as choose
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
   389
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   390
lemma gbinomial_mult_fact: "fact k * (a gchoose k) = (\<Prod>i = 0..<k. a - of_nat i)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   391
  for a :: "'a::field_char_0"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   392
  by (simp_all add: gbinomial_prod_rev field_simps)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   393
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   394
lemma gbinomial_mult_fact': "(a gchoose k) * fact k = (\<Prod>i = 0..<k. a - of_nat i)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   395
  for a :: "'a::field_char_0"
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   396
  using gbinomial_mult_fact [of k a] by (simp add: ac_simps)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   397
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   398
lemma gbinomial_pochhammer: "a gchoose k = (- 1) ^ k * pochhammer (- a) k / fact k"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   399
  for a :: "'a::field_char_0"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   400
proof (cases k)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   401
  case (Suc k')
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   402
  then have "a gchoose k = pochhammer (a - of_nat k') (Suc k') / ((1 + of_nat k') * fact k')"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   403
    by (simp add: gbinomial_prod_rev pochhammer_prod_rev atLeastLessThanSuc_atLeastAtMost
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   404
        prod.atLeast_Suc_atMost_Suc_shift of_nat_diff flip: power_mult_distrib prod.cl_ivl_Suc)
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   405
  then show ?thesis
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   406
    by (simp add: pochhammer_minus Suc)
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   407
qed auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   408
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   409
lemma gbinomial_pochhammer': "a gchoose k = pochhammer (a - of_nat k + 1) k / fact k"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   410
  for a :: "'a::field_char_0"
61552
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
   411
proof -
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   412
  have "a gchoose k = ((-1)^k * (-1)^k) * pochhammer (a - of_nat k + 1) k / fact k"
61552
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
   413
    by (simp add: gbinomial_pochhammer pochhammer_minus mult_ac)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   414
  also have "(-1 :: 'a)^k * (-1)^k = 1"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   415
    by (subst power_add [symmetric]) simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   416
  finally show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   417
    by simp
61552
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
   418
qed
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 61531
diff changeset
   419
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   420
lemma gbinomial_binomial: "n gchoose k = n choose k"
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   421
proof (cases "k \<le> n")
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   422
  case False
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   423
  then have "n < k"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   424
    by (simp add: not_le)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67299
diff changeset
   425
  then have "0 \<in> ((-) n) ` {0..<k}"
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   426
    by auto
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67299
diff changeset
   427
  then have "prod ((-) n) {0..<k} = 0"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   428
    by (auto intro: prod_zero)
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   429
  with \<open>n < k\<close> show ?thesis
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   430
    by (simp add: binomial_eq_0 gbinomial_prod_rev prod_zero)
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   431
next
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   432
  case True
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67299
diff changeset
   433
  from True have *: "prod ((-) n) {0..<k} = \<Prod>{Suc (n - k)..n}"
65350
b149abe619f7 added shuffle product to HOL/List
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   434
    by (intro prod.reindex_bij_witness[of _ "\<lambda>i. n - i" "\<lambda>i. n - i"]) auto
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   435
  from True have "n choose k = fact n div (fact k * fact (n - k))"
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   436
    by (rule binomial_fact')
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   437
  with * show ?thesis
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   438
    by (simp add: gbinomial_prod_rev mult.commute [of "fact k"] div_mult2_eq fact_div_fact)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   439
qed
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   440
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   441
lemma of_nat_gbinomial: "of_nat (n gchoose k) = (of_nat n gchoose k :: 'a::field_char_0)"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   442
proof (cases "k \<le> n")
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   443
  case False
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   444
  then show ?thesis
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   445
    by (simp add: not_le gbinomial_binomial binomial_eq_0 gbinomial_prod_rev)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   446
next
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   447
  case True
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   448
  define m where "m = n - k"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   449
  with True have n: "n = m + k"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   450
    by arith
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   451
  from n have "fact n = ((\<Prod>i = 0..<m + k. of_nat (m + k - i) ):: 'a)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   452
    by (simp add: fact_prod_rev)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   453
  also have "\<dots> = ((\<Prod>i\<in>{0..<k} \<union> {k..<m + k}. of_nat (m + k - i)) :: 'a)"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   454
    by (simp add: ivl_disj_un)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   455
  finally have "fact n = (fact m * (\<Prod>i = 0..<k. of_nat m + of_nat k - of_nat i) :: 'a)"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   456
    using prod.shift_bounds_nat_ivl [of "\<lambda>i. of_nat (m + k - i) :: 'a" 0 k m]
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   457
    by (simp add: fact_prod_rev [of m] prod.union_disjoint of_nat_diff)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   458
  then have "fact n / fact (n - k) = ((\<Prod>i = 0..<k. of_nat n - of_nat i) :: 'a)"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   459
    by (simp add: n)
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   460
  with True have "fact k * of_nat (n gchoose k) = (fact k * (of_nat n gchoose k) :: 'a)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   461
    by (simp only: gbinomial_mult_fact [of k "of_nat n"] gbinomial_binomial [of n k] fact_binomial)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   462
  then show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   463
    by simp
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   464
qed
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   465
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   466
lemma binomial_gbinomial: "of_nat (n choose k) = (of_nat n gchoose k :: 'a::field_char_0)"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   467
  by (simp add: gbinomial_binomial [symmetric] of_nat_gbinomial)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   468
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   469
setup
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69107
diff changeset
   470
  \<open>Sign.add_const_constraint (\<^const_name>\<open>gbinomial\<close>, SOME \<^typ>\<open>'a::field_char_0 \<Rightarrow> nat \<Rightarrow> 'a\<close>)\<close>
63372
492b49535094 relating gbinomial and binomial, still using distinct definitions
haftmann
parents: 63367
diff changeset
   471
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   472
lemma gbinomial_mult_1:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   473
  fixes a :: "'a::field_char_0"
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   474
  shows "a * (a gchoose k) = of_nat k * (a gchoose k) + of_nat (Suc k) * (a gchoose (Suc k))"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   475
  (is "?l = ?r")
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   476
proof -
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   477
  have "?r = ((- 1) ^k * pochhammer (- a) k / fact k) * (of_nat k - (- a + of_nat k))"
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   478
    unfolding gbinomial_pochhammer pochhammer_Suc right_diff_distrib power_Suc
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   479
    by (auto simp add: field_simps simp del: of_nat_Suc)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   480
  also have "\<dots> = ?l"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   481
    by (simp add: field_simps gbinomial_pochhammer)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   482
  finally show ?thesis ..
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   483
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   484
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   485
lemma gbinomial_mult_1':
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   486
  "(a gchoose k) * a = of_nat k * (a gchoose k) + of_nat (Suc k) * (a gchoose (Suc k))"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   487
  for a :: "'a::field_char_0"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   488
  by (simp add: mult.commute gbinomial_mult_1)
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   489
80176
7fefa7839ac6 syntax of gchoose now the same as choose
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
   490
lemma gbinomial_Suc_Suc: "(a + 1) gchoose (Suc k) = (a gchoose k) + (a gchoose (Suc k))"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   491
  for a :: "'a::field_char_0"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   492
proof (cases k)
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   493
  case 0
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   494
  then show ?thesis by simp
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   495
next
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   496
  case (Suc h)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   497
  have eq0: "(\<Prod>i\<in>{1..k}. (a + 1) - of_nat i) = (\<Prod>i\<in>{0..h}. a - of_nat i)"
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   498
  proof (rule prod.reindex_cong)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   499
    show "{1..k} = Suc ` {0..h}"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   500
      using Suc by (auto simp add: image_Suc_atMost)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   501
  qed auto
80176
7fefa7839ac6 syntax of gchoose now the same as choose
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
   502
  have "fact (Suc k) * ((a gchoose k) + (a gchoose (Suc k))) =
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   503
      (a gchoose Suc h) * (fact (Suc (Suc h))) +
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   504
      (a gchoose Suc (Suc h)) * (fact (Suc (Suc h)))"
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63366
diff changeset
   505
    by (simp add: Suc field_simps del: fact_Suc)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   506
  also have "\<dots> =
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   507
    (a gchoose Suc h) * of_nat (Suc (Suc h) * fact (Suc h)) + (\<Prod>i=0..Suc h. a - of_nat i)"
80176
7fefa7839ac6 syntax of gchoose now the same as choose
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
   508
    by (metis atLeastLessThanSuc_atLeastAtMost fact_Suc gbinomial_mult_fact mult.commute of_nat_fact of_nat_mult)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   509
  also have "\<dots> =
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   510
    (fact (Suc h) * (a gchoose Suc h)) * of_nat (Suc (Suc h)) + (\<Prod>i=0..Suc h. a - of_nat i)"
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63366
diff changeset
   511
    by (simp only: fact_Suc mult.commute mult.left_commute of_nat_fact of_nat_id of_nat_mult)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   512
  also have "\<dots> =
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   513
    of_nat (Suc (Suc h)) * (\<Prod>i=0..h. a - of_nat i) + (\<Prod>i=0..Suc h. a - of_nat i)"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   514
    unfolding gbinomial_mult_fact atLeastLessThanSuc_atLeastAtMost by auto
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   515
  also have "\<dots> =
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   516
    (\<Prod>i=0..Suc h. a - of_nat i) + (of_nat h * (\<Prod>i=0..h. a - of_nat i) + 2 * (\<Prod>i=0..h. a - of_nat i))"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59669
diff changeset
   517
    by (simp add: field_simps)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   518
  also have "\<dots> =
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   519
    ((a gchoose Suc h) * (fact (Suc h)) * of_nat (Suc k)) + (\<Prod>i\<in>{0..Suc h}. a - of_nat i)"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   520
    unfolding gbinomial_mult_fact'
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   521
    by (simp add: comm_semiring_class.distrib field_simps Suc atLeastLessThanSuc_atLeastAtMost)
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   522
  also have "\<dots> = (\<Prod>i\<in>{0..h}. a - of_nat i) * (a + 1)"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   523
    unfolding gbinomial_mult_fact' atLeast0_atMost_Suc
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   524
    by (simp add: field_simps Suc atLeastLessThanSuc_atLeastAtMost)
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   525
  also have "\<dots> = (\<Prod>i\<in>{0..k}. (a + 1) - of_nat i)"
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   526
    using eq0
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   527
    by (simp add: Suc prod.atLeast0_atMost_Suc_shift del: prod.cl_ivl_Suc)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59669
diff changeset
   528
  also have "\<dots> = (fact (Suc k)) * ((a + 1) gchoose (Suc k))"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   529
    by (simp only: gbinomial_mult_fact atLeastLessThanSuc_atLeastAtMost)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59669
diff changeset
   530
  finally show ?thesis
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   531
    using fact_nonzero [of "Suc k"] by auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   532
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   533
80176
7fefa7839ac6 syntax of gchoose now the same as choose
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
   534
lemma gbinomial_reduce_nat: "0 < k \<Longrightarrow> a gchoose k = (a-1 gchoose k-1) + (a-1 gchoose k)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   535
  for a :: "'a::field_char_0"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59669
diff changeset
   536
  by (metis Suc_pred' diff_add_cancel gbinomial_Suc_Suc)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   537
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   538
lemma gchoose_row_sum_weighted:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   539
  "(\<Sum>k = 0..m. (r gchoose k) * (r/2 - of_nat k)) = of_nat(Suc m) / 2 * (r gchoose (Suc m))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   540
  for r :: "'a::field_char_0"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   541
  by (induct m) (simp_all add: field_simps distrib gbinomial_mult_1)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   542
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   543
lemma binomial_symmetric:
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   544
  assumes kn: "k \<le> n"
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   545
  shows "n choose k = n choose (n - k)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   546
proof -
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   547
  have kn': "n - k \<le> n"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   548
    using kn by arith
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   549
  from binomial_fact_lemma[OF kn] binomial_fact_lemma[OF kn']
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   550
  have "fact k * fact (n - k) * (n choose k) = fact (n - k) * fact (n - (n - k)) * (n choose (n - k))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   551
    by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   552
  then show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   553
    using kn by simp
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   554
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   555
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   556
lemma choose_rising_sum:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   557
  "(\<Sum>j\<le>m. ((n + j) choose n)) = ((n + m + 1) choose (n + 1))"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   558
  "(\<Sum>j\<le>m. ((n + j) choose n)) = ((n + m + 1) choose m)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   559
proof -
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   560
  show "(\<Sum>j\<le>m. ((n + j) choose n)) = ((n + m + 1) choose (n + 1))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   561
    by (induct m) simp_all
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   562
  also have "\<dots> = (n + m + 1) choose m"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   563
    by (subst binomial_symmetric) simp_all
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   564
  finally show "(\<Sum>j\<le>m. ((n + j) choose n)) = (n + m + 1) choose m" .
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   565
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   566
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   567
lemma choose_linear_sum: "(\<Sum>i\<le>n. i * (n choose i)) = n * 2 ^ (n - 1)"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   568
proof (cases n)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   569
  case 0
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   570
  then show ?thesis by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   571
next
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   572
  case (Suc m)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   573
  have "(\<Sum>i\<le>n. i * (n choose i)) = (\<Sum>i\<le>Suc m. i * (Suc m choose i))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   574
    by (simp add: Suc)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   575
  also have "\<dots> = Suc m * 2 ^ m"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   576
    unfolding sum.atMost_Suc_shift Suc_times_binomial sum_distrib_left[symmetric]
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   577
    by (simp add: choose_row_sum)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   578
  finally show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   579
    using Suc by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   580
qed
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   581
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   582
lemma choose_alternating_linear_sum:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   583
  assumes "n \<noteq> 1"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   584
  shows "(\<Sum>i\<le>n. (-1)^i * of_nat i * of_nat (n choose i) :: 'a::comm_ring_1) = 0"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   585
proof (cases n)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   586
  case 0
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   587
  then show ?thesis by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   588
next
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   589
  case (Suc m)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   590
  with assms have "m > 0"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   591
    by simp
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   592
  have "(\<Sum>i\<le>n. (-1) ^ i * of_nat i * of_nat (n choose i) :: 'a) =
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   593
      (\<Sum>i\<le>Suc m. (-1) ^ i * of_nat i * of_nat (Suc m choose i))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   594
    by (simp add: Suc)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   595
  also have "\<dots> = (\<Sum>i\<le>m. (-1) ^ (Suc i) * of_nat (Suc i * (Suc m choose Suc i)))"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   596
    by (simp only: sum.atMost_Suc_shift sum_distrib_left[symmetric] mult_ac of_nat_mult) simp
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   597
  also have "\<dots> = - of_nat (Suc m) * (\<Sum>i\<le>m. (-1) ^ i * of_nat (m choose i))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   598
    by (subst sum_distrib_left, rule sum.cong[OF refl], subst Suc_times_binomial)
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   599
       (simp add: algebra_simps)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   600
  also have "(\<Sum>i\<le>m. (-1 :: 'a) ^ i * of_nat ((m choose i))) = 0"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61738
diff changeset
   601
    using choose_alternating_sum[OF \<open>m > 0\<close>] by simp
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   602
  finally show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   603
    by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   604
qed
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   605
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   606
lemma vandermonde: "(\<Sum>k\<le>r. (m choose k) * (n choose (r - k))) = (m + n) choose r"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   607
proof (induct n arbitrary: r)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   608
  case 0
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   609
  have "(\<Sum>k\<le>r. (m choose k) * (0 choose (r - k))) = (\<Sum>k\<le>r. if k = r then (m choose k) else 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   610
    by (intro sum.cong) simp_all
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   611
  also have "\<dots> = m choose r"
68784
haftmann
parents: 68077
diff changeset
   612
    by simp
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   613
  finally show ?case
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   614
    by simp
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   615
next
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   616
  case (Suc n r)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   617
  show ?case
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   618
    by (cases r) (simp_all add: Suc [symmetric] algebra_simps sum.distrib Suc_diff_le)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   619
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   620
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   621
lemma choose_square_sum: "(\<Sum>k\<le>n. (n choose k)^2) = ((2*n) choose n)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   622
  using vandermonde[of n n n]
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   623
  by (simp add: power2_eq_square mult_2 binomial_symmetric [symmetric])
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   624
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   625
lemma pochhammer_binomial_sum:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   626
  fixes a b :: "'a::comm_ring_1"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   627
  shows "pochhammer (a + b) n = (\<Sum>k\<le>n. of_nat (n choose k) * pochhammer a k * pochhammer b (n - k))"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   628
proof (induction n arbitrary: a b)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   629
  case 0
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   630
  then show ?case by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   631
next
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   632
  case (Suc n a b)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   633
  have "(\<Sum>k\<le>Suc n. of_nat (Suc n choose k) * pochhammer a k * pochhammer b (Suc n - k)) =
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   634
      (\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a (Suc i) * pochhammer b (n - i)) +
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   635
      ((\<Sum>i\<le>n. of_nat (n choose Suc i) * pochhammer a (Suc i) * pochhammer b (n - i)) +
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   636
      pochhammer b (Suc n))"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   637
    by (subst sum.atMost_Suc_shift) (simp add: ring_distribs sum.distrib)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   638
  also have "(\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a (Suc i) * pochhammer b (n - i)) =
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   639
      a * pochhammer ((a + 1) + b) n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   640
    by (subst Suc) (simp add: sum_distrib_left pochhammer_rec mult_ac)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   641
  also have "(\<Sum>i\<le>n. of_nat (n choose Suc i) * pochhammer a (Suc i) * pochhammer b (n - i)) +
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   642
        pochhammer b (Suc n) =
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   643
      (\<Sum>i=0..Suc n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc n - i))"
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   644
    apply (subst sum.atLeast_Suc_atMost, simp)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   645
    apply (simp add: sum.shift_bounds_cl_Suc_ivl atLeast0AtMost del: sum.cl_ivl_Suc)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   646
    done
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   647
  also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc n - i))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   648
    using Suc by (intro sum.mono_neutral_right) (auto simp: not_le binomial_eq_0)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   649
  also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc (n - i)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   650
    by (intro sum.cong) (simp_all add: Suc_diff_le)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   651
  also have "\<dots> = b * pochhammer (a + (b + 1)) n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   652
    by (subst Suc) (simp add: sum_distrib_left mult_ac pochhammer_rec)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   653
  also have "a * pochhammer ((a + 1) + b) n + b * pochhammer (a + (b + 1)) n =
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   654
      pochhammer (a + b) (Suc n)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   655
    by (simp add: pochhammer_rec algebra_simps)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   656
  finally show ?case ..
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   657
qed
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   658
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   659
text \<open>Contributed by Manuel Eberl, generalised by LCP.
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69107
diff changeset
   660
  Alternative definition of the binomial coefficient as \<^term>\<open>\<Prod>i<k. (n - i) / (k - i)\<close>.\<close>
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   661
lemma gbinomial_altdef_of_nat: "a gchoose k = (\<Prod>i = 0..<k. (a - of_nat i) / of_nat (k - i) :: 'a)"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   662
  for k :: nat and a :: "'a::field_char_0"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   663
  by (simp add: prod_dividef gbinomial_prod_rev fact_prod_rev)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   664
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   665
lemma gbinomial_ge_n_over_k_pow_k:
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   666
  fixes k :: nat
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   667
    and a :: "'a::linordered_field"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   668
  assumes "of_nat k \<le> a"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   669
  shows "(a / of_nat k :: 'a) ^ k \<le> a gchoose k"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   670
proof -
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   671
  have x: "0 \<le> a"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   672
    using assms of_nat_0_le_iff order_trans by blast
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   673
  have "(a / of_nat k :: 'a) ^ k = (\<Prod>i = 0..<k. a / of_nat k :: 'a)"
68784
haftmann
parents: 68077
diff changeset
   674
    by simp
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   675
  also have "\<dots> \<le> a gchoose k"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   676
  proof -
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   677
    have "\<And>i. i < k \<Longrightarrow> 0 \<le> a / of_nat k"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   678
      by (simp add: x zero_le_divide_iff)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   679
    moreover have "a / of_nat k \<le> (a - of_nat i) / of_nat (k - i)" if "i < k" for i
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   680
    proof -
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   681
      from assms have "a * of_nat i \<ge> of_nat (i * k)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   682
        by (metis mult.commute mult_le_cancel_right of_nat_less_0_iff of_nat_mult)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   683
      then have "a * of_nat k - a * of_nat i \<le> a * of_nat k - of_nat (i * k)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   684
        by arith
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   685
      then have "a * of_nat (k - i) \<le> (a - of_nat i) * of_nat k"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   686
        using \<open>i < k\<close> by (simp add: algebra_simps zero_less_mult_iff of_nat_diff)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   687
      then have "a * of_nat (k - i) \<le> (a - of_nat i) * (of_nat k :: 'a)"
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   688
        by blast
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   689
      with assms show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   690
        using \<open>i < k\<close> by (simp add: field_simps)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   691
    qed
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   692
    ultimately show ?thesis
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   693
      unfolding gbinomial_altdef_of_nat
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   694
      by (intro prod_mono) auto
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   695
  qed
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   696
  finally show ?thesis .
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   697
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
   698
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   699
lemma gbinomial_negated_upper: "(a gchoose k) = (-1) ^ k * ((of_nat k - a - 1) gchoose k)"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   700
  by (simp add: gbinomial_pochhammer pochhammer_minus algebra_simps)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   701
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   702
lemma gbinomial_minus: "((-a) gchoose k) = (-1) ^ k * ((a + of_nat k - 1) gchoose k)"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   703
  by (subst gbinomial_negated_upper) (simp add: add_ac)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   704
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   705
lemma Suc_times_gbinomial: "of_nat (Suc k) * ((a + 1) gchoose (Suc k)) = (a + 1) * (a gchoose k)"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   706
proof (cases k)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   707
  case 0
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   708
  then show ?thesis by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   709
next
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   710
  case (Suc b)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   711
  then have "((a + 1) gchoose (Suc (Suc b))) = (\<Prod>i = 0..Suc b. a + (1 - of_nat i)) / fact (b + 2)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   712
    by (simp add: field_simps gbinomial_prod_rev atLeastLessThanSuc_atLeastAtMost)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   713
  also have "(\<Prod>i = 0..Suc b. a + (1 - of_nat i)) = (a + 1) * (\<Prod>i = 0..b. a - of_nat i)"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   714
    by (simp add: prod.atLeast0_atMost_Suc_shift del: prod.cl_ivl_Suc)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   715
  also have "\<dots> / fact (b + 2) = (a + 1) / of_nat (Suc (Suc b)) * (a gchoose Suc b)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   716
    by (simp_all add: gbinomial_prod_rev atLeastLessThanSuc_atLeastAtMost)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   717
  finally show ?thesis by (simp add: Suc field_simps del: of_nat_Suc)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   718
qed
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   719
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   720
lemma gbinomial_factors: "((a + 1) gchoose (Suc k)) = (a + 1) / of_nat (Suc k) * (a gchoose k)"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   721
proof (cases k)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   722
  case 0
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   723
  then show ?thesis by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   724
next
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   725
  case (Suc b)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   726
  then have "((a + 1) gchoose (Suc (Suc b))) = (\<Prod>i = 0 .. Suc b. a + (1 - of_nat i)) / fact (b + 2)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   727
    by (simp add: field_simps gbinomial_prod_rev atLeastLessThanSuc_atLeastAtMost)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   728
  also have "(\<Prod>i = 0 .. Suc b. a + (1 - of_nat i)) = (a + 1) * (\<Prod>i = 0..b. a - of_nat i)"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   729
    by (simp add: prod.atLeast0_atMost_Suc_shift del: prod.cl_ivl_Suc)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   730
  also have "\<dots> / fact (b + 2) = (a + 1) / of_nat (Suc (Suc b)) * (a gchoose Suc b)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   731
    by (simp_all add: gbinomial_prod_rev atLeastLessThanSuc_atLeastAtMost atLeast0AtMost)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   732
  finally show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   733
    by (simp add: Suc)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   734
qed
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   735
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   736
lemma gbinomial_rec: "((a + 1) gchoose (Suc k)) = (a gchoose k) * ((a + 1) / of_nat (Suc k))"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   737
  using gbinomial_mult_1[of a k]
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   738
  by (subst gbinomial_Suc_Suc) (simp add: field_simps del: of_nat_Suc, simp add: algebra_simps)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   739
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   740
lemma gbinomial_of_nat_symmetric: "k \<le> n \<Longrightarrow> (of_nat n) gchoose k = (of_nat n) gchoose (n - k)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   741
  using binomial_symmetric[of k n] by (simp add: binomial_gbinomial [symmetric])
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   742
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   743
77172
816959264c32 isabelle update -u cite -l "";
wenzelm
parents: 75865
diff changeset
   744
text \<open>The absorption identity (equation 5.5 \<^cite>\<open>\<open>p.~157\<close> in GKP_CM\<close>):
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   745
\[
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   746
{r \choose k} = \frac{r}{k}{r - 1 \choose k - 1},\quad \textnormal{integer } k \neq 0.
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   747
\]\<close>
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   748
lemma gbinomial_absorption': "k > 0 \<Longrightarrow> a gchoose k = (a / of_nat k) * (a - 1 gchoose (k - 1))"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   749
  using gbinomial_rec[of "a - 1" "k - 1"]
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   750
  by (simp_all add: gbinomial_rec field_simps del: of_nat_Suc)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   751
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   752
text \<open>The absorption identity is written in the following form to avoid
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   753
division by $k$ (the lower index) and therefore remove the $k \neq 0$
77172
816959264c32 isabelle update -u cite -l "";
wenzelm
parents: 75865
diff changeset
   754
restriction \<^cite>\<open>\<open>p.~157\<close> in GKP_CM\<close>:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   755
\[
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   756
k{r \choose k} = r{r - 1 \choose k - 1}, \quad \textnormal{integer } k.
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   757
\]\<close>
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   758
lemma gbinomial_absorption: "of_nat (Suc k) * (a gchoose Suc k) = a * ((a - 1) gchoose k)"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   759
  using gbinomial_absorption'[of "Suc k" a] by (simp add: field_simps del: of_nat_Suc)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   760
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   761
text \<open>The absorption identity for natural number binomial coefficients:\<close>
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   762
lemma binomial_absorption: "Suc k * (n choose Suc k) = n * ((n - 1) choose k)"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   763
  by (cases n) (simp_all add: binomial_eq_0 Suc_times_binomial del: binomial_Suc_Suc mult_Suc)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   764
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   765
text \<open>The absorption companion identity for natural number coefficients,
77172
816959264c32 isabelle update -u cite -l "";
wenzelm
parents: 75865
diff changeset
   766
  following the proof by GKP \<^cite>\<open>\<open>p.~157\<close> in GKP_CM\<close>:\<close>
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   767
lemma binomial_absorb_comp: "(n - k) * (n choose k) = n * ((n - 1) choose k)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   768
  (is "?lhs = ?rhs")
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   769
proof (cases "n \<le> k")
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   770
  case True
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   771
  then show ?thesis by auto
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   772
next
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   773
  case False
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   774
  then have "?rhs = Suc ((n - 1) - k) * (n choose Suc ((n - 1) - k))"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   775
    using binomial_symmetric[of k "n - 1"] binomial_absorption[of "(n - 1) - k" n]
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   776
    by simp
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   777
  also have "Suc ((n - 1) - k) = n - k"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   778
    using False by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   779
  also have "n choose \<dots> = n choose k"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   780
    using False by (intro binomial_symmetric [symmetric]) simp_all
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   781
  finally show ?thesis ..
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   782
qed
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   783
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   784
text \<open>The generalised absorption companion identity:\<close>
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   785
lemma gbinomial_absorb_comp: "(a - of_nat k) * (a gchoose k) = a * ((a - 1) gchoose k)"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   786
  using pochhammer_absorb_comp[of a k] by (simp add: gbinomial_pochhammer)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   787
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   788
lemma gbinomial_addition_formula:
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   789
  "a gchoose (Suc k) = ((a - 1) gchoose (Suc k)) + ((a - 1) gchoose k)"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   790
  using gbinomial_Suc_Suc[of "a - 1" k] by (simp add: algebra_simps)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   791
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   792
lemma binomial_addition_formula:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   793
  "0 < n \<Longrightarrow> n choose (Suc k) = ((n - 1) choose (Suc k)) + ((n - 1) choose k)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   794
  by (subst choose_reduce_nat) simp_all
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   795
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   796
text \<open>
77172
816959264c32 isabelle update -u cite -l "";
wenzelm
parents: 75865
diff changeset
   797
  Equation 5.9 of the reference material \<^cite>\<open>\<open>p.~159\<close> in GKP_CM\<close> is a useful
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   798
  summation formula, operating on both indices:
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   799
  \[
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   800
   \sum\limits_{k \leq n}{r + k \choose k} = {r + n + 1 \choose n},
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   801
   \quad \textnormal{integer } n.
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   802
  \]
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   803
\<close>
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   804
lemma gbinomial_parallel_sum: "(\<Sum>k\<le>n. (a + of_nat k) gchoose k) = (a + of_nat n + 1) gchoose n"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   805
proof (induct n)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   806
  case 0
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   807
  then show ?case by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   808
next
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   809
  case (Suc m)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   810
  then show ?case
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   811
    using gbinomial_Suc_Suc[of "(a + of_nat m + 1)" m]
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   812
    by (simp add: add_ac)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   813
qed
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   814
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   815
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
   816
subsection \<open>Summation on the upper index\<close>
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   817
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   818
text \<open>
77172
816959264c32 isabelle update -u cite -l "";
wenzelm
parents: 75865
diff changeset
   819
  Another summation formula is equation 5.10 of the reference material \<^cite>\<open>\<open>p.~160\<close> in GKP_CM\<close>,
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   820
  aptly named \emph{summation on the upper index}:\[\sum_{0 \leq k \leq n} {k \choose m} =
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   821
  {n + 1 \choose m + 1}, \quad \textnormal{integers } m, n \geq 0.\]
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   822
\<close>
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   823
lemma gbinomial_sum_up_index:
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   824
  "(\<Sum>j = 0..n. (of_nat j gchoose k) :: 'a::field_char_0) = (of_nat n + 1) gchoose (k + 1)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   825
proof (induct n)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   826
  case 0
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   827
  show ?case
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   828
    using gbinomial_Suc_Suc[of 0 k]
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   829
    by (cases k) auto
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   830
next
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   831
  case (Suc n)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   832
  then show ?case
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   833
    using gbinomial_Suc_Suc[of "of_nat (Suc n) :: 'a" k]
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   834
    by (simp add: add_ac)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   835
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   836
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   837
lemma gbinomial_index_swap:
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   838
  "((-1) ^ k) * ((- (of_nat n) - 1) gchoose k) = ((-1) ^ n) * ((- (of_nat k) - 1) gchoose n)"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   839
  (is "?lhs = ?rhs")
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   840
proof -
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   841
  have "?lhs = (of_nat (k + n) gchoose k)"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   842
    by (subst gbinomial_negated_upper) (simp add: power_mult_distrib [symmetric])
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   843
  also have "\<dots> = (of_nat (k + n) gchoose n)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   844
    by (subst gbinomial_of_nat_symmetric) simp_all
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   845
  also have "\<dots> = ?rhs"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   846
    by (subst gbinomial_negated_upper) simp
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   847
  finally show ?thesis .
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   848
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   849
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   850
lemma gbinomial_sum_lower_neg: "(\<Sum>k\<le>m. (a gchoose k) * (- 1) ^ k) = (- 1) ^ m * (a - 1 gchoose m)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   851
  (is "?lhs = ?rhs")
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   852
proof -
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   853
  have "?lhs = (\<Sum>k\<le>m. -(a + 1) + of_nat k gchoose k)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   854
    by (intro sum.cong[OF refl]) (subst gbinomial_negated_upper, simp add: power_mult_distrib)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   855
  also have "\<dots>  = - a + of_nat m gchoose m"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   856
    by (subst gbinomial_parallel_sum) simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   857
  also have "\<dots> = ?rhs"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   858
    by (subst gbinomial_negated_upper) (simp add: power_mult_distrib)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   859
  finally show ?thesis .
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   860
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   861
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   862
lemma gbinomial_partial_row_sum:
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   863
  "(\<Sum>k\<le>m. (a gchoose k) * ((a / 2) - of_nat k)) = ((of_nat m + 1)/2) * (a gchoose (m + 1))"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   864
proof (induct m)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   865
  case 0
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   866
  then show ?case by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   867
next
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   868
  case (Suc mm)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   869
  then have "(\<Sum>k\<le>Suc mm. (a gchoose k) * (a / 2 - of_nat k)) =
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   870
      (a - of_nat (Suc mm)) * (a gchoose Suc mm) / 2"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   871
    by (simp add: field_simps)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   872
  also have "\<dots> = a * (a - 1 gchoose Suc mm) / 2"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   873
    by (subst gbinomial_absorb_comp) (rule refl)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   874
  also have "\<dots> = (of_nat (Suc mm) + 1) / 2 * (a gchoose (Suc mm + 1))"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   875
    by (subst gbinomial_absorption [symmetric]) simp
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   876
  finally show ?case .
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   877
qed
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   878
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   879
lemma sum_bounds_lt_plus1: "(\<Sum>k<mm. f (Suc k)) = (\<Sum>k=1..mm. f k)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   880
  by (induct mm) simp_all
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   881
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   882
lemma gbinomial_partial_sum_poly:
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   883
  "(\<Sum>k\<le>m. (of_nat m + a gchoose k) * x^k * y^(m-k)) =
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   884
    (\<Sum>k\<le>m. (-a gchoose k) * (-x)^k * (x + y)^(m-k))"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   885
  (is "?lhs m = ?rhs m")
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   886
proof (induction m)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   887
  case 0
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   888
  then show ?case by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   889
next
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   890
  case (Suc mm)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   891
  define G where "G i k = (of_nat i + a gchoose k) * x^k * y^(i - k)" for i k
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62481
diff changeset
   892
  define S where "S = ?lhs"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   893
  have SG_def: "S = (\<lambda>i. (\<Sum>k\<le>i. (G i k)))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   894
    unfolding S_def G_def ..
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   895
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   896
  have "S (Suc mm) = G (Suc mm) 0 + (\<Sum>k=Suc 0..Suc mm. G (Suc mm) k)"
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69768
diff changeset
   897
    using SG_def by (simp add: sum.atLeast_Suc_atMost atLeast0AtMost [symmetric])
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   898
  also have "(\<Sum>k=Suc 0..Suc mm. G (Suc mm) k) = (\<Sum>k=0..mm. G (Suc mm) (Suc k))"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   899
    by (subst sum.shift_bounds_cl_Suc_ivl) simp
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   900
  also have "\<dots> = (\<Sum>k=0..mm. ((of_nat mm + a gchoose (Suc k)) +
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   901
      (of_nat mm + a gchoose k)) * x^(Suc k) * y^(mm - k))"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   902
    unfolding G_def by (subst gbinomial_addition_formula) simp
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   903
  also have "\<dots> = (\<Sum>k=0..mm. (of_nat mm + a gchoose (Suc k)) * x^(Suc k) * y^(mm - k)) +
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   904
      (\<Sum>k=0..mm. (of_nat mm + a gchoose k) * x^(Suc k) * y^(mm - k))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   905
    by (subst sum.distrib [symmetric]) (simp add: algebra_simps)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   906
  also have "(\<Sum>k=0..mm. (of_nat mm + a gchoose (Suc k)) * x^(Suc k) * y^(mm - k)) =
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   907
      (\<Sum>k<Suc mm. (of_nat mm + a gchoose (Suc k)) * x^(Suc k) * y^(mm - k))"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   908
    by (simp only: atLeast0AtMost lessThan_Suc_atMost)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   909
  also have "\<dots> = (\<Sum>k<mm. (of_nat mm + a gchoose Suc k) * x^(Suc k) * y^(mm-k)) +
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   910
      (of_nat mm + a gchoose (Suc mm)) * x^(Suc mm)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   911
    (is "_ = ?A + ?B")
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69768
diff changeset
   912
    by (subst sum.lessThan_Suc) simp
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   913
  also have "?A = (\<Sum>k=1..mm. (of_nat mm + a gchoose k) * x^k * y^(mm - k + 1))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   914
  proof (subst sum_bounds_lt_plus1 [symmetric], intro sum.cong[OF refl], clarify)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   915
    fix k
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   916
    assume "k < mm"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   917
    then have "mm - k = mm - Suc k + 1"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   918
      by linarith
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   919
    then show "(of_nat mm + a gchoose Suc k) * x ^ Suc k * y ^ (mm - k) =
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   920
        (of_nat mm + a gchoose Suc k) * x ^ Suc k * y ^ (mm - Suc k + 1)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   921
      by (simp only:)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   922
  qed
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   923
  also have "\<dots> + ?B = y * (\<Sum>k=1..mm. (G mm k)) + (of_nat mm + a gchoose (Suc mm)) * x^(Suc mm)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   924
    unfolding G_def by (subst sum_distrib_left) (simp add: algebra_simps)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   925
  also have "(\<Sum>k=0..mm. (of_nat mm + a gchoose k) * x^(Suc k) * y^(mm - k)) = x * (S mm)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   926
    unfolding S_def by (subst sum_distrib_left) (simp add: atLeast0AtMost algebra_simps)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   927
  also have "(G (Suc mm) 0) = y * (G mm 0)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   928
    by (simp add: G_def)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   929
  finally have "S (Suc mm) =
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   930
      y * (G mm 0 + (\<Sum>k=1..mm. (G mm k))) + (of_nat mm + a gchoose (Suc mm)) * x^(Suc mm) + x * (S mm)"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   931
    by (simp add: ring_distribs)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   932
  also have "G mm 0 + (\<Sum>k=1..mm. (G mm k)) = S mm"
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69768
diff changeset
   933
    by (simp add: sum.atLeast_Suc_atMost[symmetric] SG_def atLeast0AtMost)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   934
  finally have "S (Suc mm) = (x + y) * (S mm) + (of_nat mm + a gchoose (Suc mm)) * x^(Suc mm)"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   935
    by (simp add: algebra_simps)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   936
  also have "(of_nat mm + a gchoose (Suc mm)) = (-1) ^ (Suc mm) * (- a gchoose (Suc mm))"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   937
    by (subst gbinomial_negated_upper) simp
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   938
  also have "(-1) ^ Suc mm * (- a gchoose Suc mm) * x ^ Suc mm =
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   939
      (- a gchoose (Suc mm)) * (-x) ^ Suc mm"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   940
    by (simp add: power_minus[of x])
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   941
  also have "(x + y) * S mm + \<dots> = (x + y) * ?rhs mm + (- a gchoose (Suc mm)) * (- x)^Suc mm"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   942
    unfolding S_def by (subst Suc.IH) simp
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   943
  also have "(x + y) * ?rhs mm = (\<Sum>n\<le>mm. ((- a gchoose n) * (- x) ^ n * (x + y) ^ (Suc mm - n)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   944
    by (subst sum_distrib_left, rule sum.cong) (simp_all add: Suc_diff_le)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   945
  also have "\<dots> + (-a gchoose (Suc mm)) * (-x)^Suc mm =
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   946
      (\<Sum>n\<le>Suc mm. (- a gchoose n) * (- x) ^ n * (x + y) ^ (Suc mm - n))"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   947
    by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   948
  finally show ?case
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   949
    by (simp only: S_def)
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   950
qed
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   951
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   952
lemma gbinomial_partial_sum_poly_xpos:
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
   953
    "(\<Sum>k\<le>m. (of_nat m + a gchoose k) * x^k * y^(m-k)) =
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   954
     (\<Sum>k\<le>m. (of_nat k + a - 1 gchoose k) * x^k * (x + y)^(m-k))" (is "?lhs = ?rhs")
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   955
proof -
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   956
  have "?lhs = (\<Sum>k\<le>m. (- a gchoose k) * (- x) ^ k * (x + y) ^ (m - k))"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   957
    by (simp add: gbinomial_partial_sum_poly)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   958
  also have "... = (\<Sum>k\<le>m. (-1) ^ k * (of_nat k - - a - 1 gchoose k) * (- x) ^ k * (x + y) ^ (m - k))"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 72302
diff changeset
   959
    by (metis (no_types, opaque_lifting) gbinomial_negated_upper)
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   960
  also have "... = ?rhs"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   961
    by (intro sum.cong) (auto simp flip: power_mult_distrib)
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   962
  finally show ?thesis .
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
   963
qed
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   964
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   965
lemma binomial_r_part_sum: "(\<Sum>k\<le>m. (2 * m + 1 choose k)) = 2 ^ (2 * m)"
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   966
proof -
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   967
  have "2 * 2^(2*m) = (\<Sum>k = 0..(2 * m + 1). (2 * m + 1 choose k))"
68077
ee8c13ae81e9 Some tidying up (mostly regarding summations from 0)
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   968
    using choose_row_sum[where n="2 * m + 1"]  by (simp add: atMost_atLeast0)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   969
  also have "(\<Sum>k = 0..(2 * m + 1). (2 * m + 1 choose k)) =
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   970
      (\<Sum>k = 0..m. (2 * m + 1 choose k)) +
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   971
      (\<Sum>k = m+1..2*m+1. (2 * m + 1 choose k))"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   972
    using sum.ub_add_nat[of 0 m "\<lambda>k. 2 * m + 1 choose k" "m+1"]
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   973
    by (simp add: mult_2)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   974
  also have "(\<Sum>k = m+1..2*m+1. (2 * m + 1 choose k)) =
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   975
      (\<Sum>k = 0..m. (2 * m + 1 choose (k + (m + 1))))"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
   976
    by (subst sum.shift_bounds_cl_nat_ivl [symmetric]) (simp add: mult_2)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   977
  also have "\<dots> = (\<Sum>k = 0..m. (2 * m + 1 choose (m - k)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   978
    by (intro sum.cong[OF refl], subst binomial_symmetric) simp_all
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   979
  also have "\<dots> = (\<Sum>k = 0..m. (2 * m + 1 choose k))"
67411
3f4b0c84630f restored naming of lemmas after corresponding constants
haftmann
parents: 67399
diff changeset
   980
    using sum.atLeastAtMost_rev [of "\<lambda>k. 2 * m + 1 choose (m - k)" 0 m]
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63373
diff changeset
   981
    by simp
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   982
  also have "\<dots> + \<dots> = 2 * \<dots>"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   983
    by simp
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   984
  finally show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   985
    by (subst (asm) mult_cancel1) (simp add: atLeast0AtMost)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   986
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   987
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   988
lemma gbinomial_r_part_sum: "(\<Sum>k\<le>m. (2 * (of_nat m) + 1 gchoose k)) = 2 ^ (2 * m)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   989
  (is "?lhs = ?rhs")
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   990
proof -
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
   991
  have "?lhs = of_nat (\<Sum>k\<le>m. (2 * m + 1) choose k)"
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   992
    by (simp add: binomial_gbinomial add_ac)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   993
  also have "\<dots> = of_nat (2 ^ (2 * m))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   994
    by (subst binomial_r_part_sum) (rule refl)
63366
209c4cbbc4cd define binomial coefficents directly via combinatorial definition
haftmann
parents: 63363
diff changeset
   995
  finally show ?thesis by simp
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   996
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   997
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
   998
lemma gbinomial_sum_nat_pow2:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
   999
  "(\<Sum>k\<le>m. (of_nat (m + k) gchoose k :: 'a::field_char_0) / 2 ^ k) = 2 ^ m"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1000
  (is "?lhs = ?rhs")
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1001
proof -
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1002
  have "2 ^ m * 2 ^ m = (2 ^ (2*m) :: 'a)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1003
    by (induct m) simp_all
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1004
  also have "\<dots> = (\<Sum>k\<le>m. (2 * (of_nat m) + 1 gchoose k))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1005
    using gbinomial_r_part_sum ..
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1006
  also have "\<dots> = (\<Sum>k\<le>m. (of_nat (m + k) gchoose k) * 2 ^ (m - k))"
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
  1007
    using gbinomial_partial_sum_poly_xpos[where x="1" and y="1" and a="of_nat m + 1" and m="m"]
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1008
    by (simp add: add_ac)
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1009
  also have "\<dots> = 2 ^ m * (\<Sum>k\<le>m. (of_nat (m + k) gchoose k) / 2 ^ k)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  1010
    by (subst sum_distrib_left) (simp add: algebra_simps power_diff)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1011
  finally show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1012
    by (subst (asm) mult_left_cancel) simp_all
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1013
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1014
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1015
lemma gbinomial_trinomial_revision:
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1016
  assumes "k \<le> m"
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
  1017
  shows "(a gchoose m) * (of_nat m gchoose k) = (a gchoose k) * (a - of_nat k gchoose (m - k))"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1018
proof -
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
  1019
  have "(a gchoose m) * (of_nat m gchoose k) = (a gchoose m) * fact m / (fact k * fact (m - k))"
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1020
    using assms by (simp add: binomial_gbinomial [symmetric] binomial_fact)
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
  1021
  also have "\<dots> = (a gchoose k) * (a - of_nat k gchoose (m - k))"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1022
    using assms by (simp add: gbinomial_pochhammer power_diff pochhammer_product)
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1023
  finally show ?thesis .
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1024
qed
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61076
diff changeset
  1025
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1026
text \<open>Versions of the theorems above for the natural-number version of "choose"\<close>
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1027
lemma binomial_altdef_of_nat:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1028
  "k \<le> n \<Longrightarrow> of_nat (n choose k) = (\<Prod>i = 0..<k. of_nat (n - i) / of_nat (k - i) :: 'a)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1029
  for n k :: nat and x :: "'a::field_char_0"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1030
  by (simp add: gbinomial_altdef_of_nat binomial_gbinomial of_nat_diff)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1031
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1032
lemma binomial_ge_n_over_k_pow_k: "k \<le> n \<Longrightarrow> (of_nat n / of_nat k :: 'a) ^ k \<le> of_nat (n choose k)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1033
  for k n :: nat and x :: "'a::linordered_field"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1034
  by (simp add: gbinomial_ge_n_over_k_pow_k binomial_gbinomial of_nat_diff)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1035
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1036
lemma binomial_le_pow:
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1037
  assumes "r \<le> n"
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1038
  shows "n choose r \<le> n ^ r"
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1039
proof -
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1040
  have "n choose r \<le> fact n div fact (n - r)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1041
    using assms by (subst binomial_fact_lemma[symmetric]) auto
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1042
  with fact_div_fact_le_pow [OF assms] show ?thesis
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1043
    by auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1044
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1045
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1046
lemma binomial_altdef_nat: "k \<le> n \<Longrightarrow> n choose k = fact n div (fact k * fact (n - k))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1047
  for k n :: nat
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1048
  by (subst binomial_fact_lemma [symmetric]) auto
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1049
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1050
lemma choose_dvd:
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
  1051
  assumes "k \<le> n" shows "fact k * fact (n - k) dvd (fact n :: 'a::linordered_semidom)"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59669
diff changeset
  1052
  unfolding dvd_def
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
  1053
proof
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
  1054
  show "fact n = fact k * fact (n - k) * of_nat (n choose k)"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
  1055
    by (metis assms binomial_fact_lemma of_nat_fact of_nat_mult) 
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
  1056
qed
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1057
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
  1058
lemma fact_fact_dvd_fact:
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66311
diff changeset
  1059
  "fact k * fact n dvd (fact (k + n) :: 'a::linordered_semidom)"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1060
  by (metis add.commute add_diff_cancel_left' choose_dvd le_add2)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1061
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1062
lemma choose_mult_lemma:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1063
  "((m + r + k) choose (m + k)) * ((m + k) choose k) = ((m + r + k) choose k) * ((m + r) choose m)"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1064
  (is "?lhs = _")
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1065
proof -
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1066
  have "?lhs =
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1067
      fact (m + r + k) div (fact (m + k) * fact (m + r - m)) * (fact (m + k) div (fact k * fact m))"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1068
    by (simp add: binomial_altdef_nat)
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
  1069
  also have "... = fact (m + r + k) * fact (m + k) div
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
  1070
                 (fact (m + k) * fact (m + r - m) * (fact k * fact m))"
75864
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1071
    by (metis add_implies_diff add_le_mono1 choose_dvd diff_cancel2 div_mult_div_if_dvd le_add1 le_add2)
71699
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
  1072
  also have "\<dots> = fact (m + r + k) div (fact r * (fact k * fact m))"
8e5c20e4e11a a few more applys
paulson <lp15@cam.ac.uk>
parents: 71351
diff changeset
  1073
    by (auto simp: algebra_simps fact_fact_dvd_fact)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1074
  also have "\<dots> = (fact (m + r + k) * fact (m + r)) div (fact r * (fact k * fact m) * fact (m + r))"
75864
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1075
    by simp
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1076
  also have "\<dots> =
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1077
      (fact (m + r + k) div (fact k * fact (m + r)) * (fact (m + r) div (fact r * fact m)))"
71720
1d8a1f727879 removed more applys
paulson <lp15@cam.ac.uk>
parents: 71699
diff changeset
  1078
    by (auto simp: div_mult_div_if_dvd fact_fact_dvd_fact algebra_simps)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1079
  finally show ?thesis
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1080
    by (simp add: binomial_altdef_nat mult.commute)
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1081
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1082
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1083
text \<open>The "Subset of a Subset" identity.\<close>
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1084
lemma choose_mult:
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1085
  "k \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> (n choose m) * (m choose k) = (n choose k) * ((n - k) choose (m - k))"
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1086
  using choose_mult_lemma [of "m-k" "n-m" k] by simp
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1087
75864
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1088
lemma of_nat_binomial_eq_mult_binomial_Suc:
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1089
  assumes "k \<le> n"
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1090
  shows "(of_nat :: (nat \<Rightarrow> ('a :: field_char_0))) (n choose k) = of_nat (n + 1 - k) / of_nat (n + 1) * of_nat (Suc n choose k)"
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1091
proof (cases k)
75865
62c64e3e4741 The same, without adding a new simprule
paulson <lp15@cam.ac.uk>
parents: 75864
diff changeset
  1092
  case 0 then show ?thesis
62c64e3e4741 The same, without adding a new simprule
paulson <lp15@cam.ac.uk>
parents: 75864
diff changeset
  1093
    using of_nat_neq_0 by auto
75864
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1094
next
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1095
  case (Suc l)
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1096
  have "of_nat (n + 1) * (\<Prod>i=0..<k. of_nat (n - i)) = (of_nat :: (nat \<Rightarrow> 'a)) (n + 1 - k) * (\<Prod>i=0..<k. of_nat (Suc n - i))"
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1097
    using prod.atLeast0_lessThan_Suc [where ?'a = 'a, symmetric, of "\<lambda>i. of_nat (Suc n - i)" k]
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1098
    by (simp add: ac_simps prod.atLeast0_lessThan_Suc_shift del: prod.op_ivl_Suc)
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1099
  also have "... = (of_nat :: (nat \<Rightarrow> 'a)) (Suc n - k) * (\<Prod>i=0..<k. of_nat (Suc n - i))"
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1100
    by (simp add: Suc atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost)
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1101
  also have "... = (of_nat :: (nat \<Rightarrow> 'a)) (n + 1 - k) * (\<Prod>i=0..<k. of_nat (Suc n - i))"
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1102
    by (simp only: Suc_eq_plus1)
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1103
  finally have "(\<Prod>i=0..<k. of_nat (n - i)) = (of_nat :: (nat \<Rightarrow> 'a)) (n + 1 - k) / of_nat (n + 1) * (\<Prod>i=0..<k. of_nat (Suc n - i))"
75865
62c64e3e4741 The same, without adding a new simprule
paulson <lp15@cam.ac.uk>
parents: 75864
diff changeset
  1104
    using of_nat_neq_0 by (auto simp: mult.commute divide_simps)
75864
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1105
  with assms show ?thesis
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1106
    by (simp add: binomial_altdef_of_nat prod_dividef)
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1107
qed
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 75856
diff changeset
  1108
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1109
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1110
subsection \<open>More on Binomial Coefficients\<close>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1111
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1112
text \<open>The number of nat lists of length \<open>m\<close> summing to \<open>N\<close> is \<^term>\<open>(N + m - 1) choose N\<close>:\<close>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1113
lemma card_length_sum_list_rec:
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1114
  assumes "m \<ge> 1"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1115
  shows "card {l::nat list. length l = m \<and> sum_list l = N} =
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1116
      card {l. length l = (m - 1) \<and> sum_list l = N} +
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1117
      card {l. length l = m \<and> sum_list l + 1 = N}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1118
    (is "card ?C = card ?A + card ?B")
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1119
proof -
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1120
  let ?A' = "{l. length l = m \<and> sum_list l = N \<and> hd l = 0}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1121
  let ?B' = "{l. length l = m \<and> sum_list l = N \<and> hd l \<noteq> 0}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1122
  let ?f = "\<lambda>l. 0 # l"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1123
  let ?g = "\<lambda>l. (hd l + 1) # tl l"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1124
  have 1: "xs \<noteq> [] \<Longrightarrow> x = hd xs \<Longrightarrow> x # tl xs = xs" for x :: nat and xs
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1125
    by simp
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1126
  have 2: "xs \<noteq> [] \<Longrightarrow> sum_list(tl xs) = sum_list xs - hd xs" for xs :: "nat list"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1127
    by (auto simp add: neq_Nil_conv)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1128
  have f: "bij_betw ?f ?A ?A'"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1129
    by (rule bij_betw_byWitness[where f' = tl]) (use assms in \<open>auto simp: 2 1 simp flip: length_0_conv\<close>)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1130
  have 3: "xs \<noteq> [] \<Longrightarrow> hd xs + (sum_list xs - hd xs) = sum_list xs" for xs :: "nat list"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1131
    by (metis 1 sum_list_simps(2) 2)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1132
  have g: "bij_betw ?g ?B ?B'"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1133
    apply (rule bij_betw_byWitness[where f' = "\<lambda>l. (hd l - 1) # tl l"])
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1134
    using assms
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1135
    by (auto simp: 2 simp flip: length_0_conv intro!: 3)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1136
  have fin: "finite {xs. size xs = M \<and> set xs \<subseteq> {0..<N}}" for M N :: nat
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1137
    using finite_lists_length_eq[OF finite_atLeastLessThan] conj_commute by auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1138
  have fin_A: "finite ?A" using fin[of _ "N+1"]
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1139
    by (intro finite_subset[where ?A = "?A" and ?B = "{xs. size xs = m - 1 \<and> set xs \<subseteq> {0..<N+1}}"])
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1140
      (auto simp: member_le_sum_list less_Suc_eq_le)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1141
  have fin_B: "finite ?B"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1142
    by (intro finite_subset[where ?A = "?B" and ?B = "{xs. size xs = m \<and> set xs \<subseteq> {0..<N}}"])
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1143
      (auto simp: member_le_sum_list less_Suc_eq_le fin)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1144
  have uni: "?C = ?A' \<union> ?B'"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1145
    by auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1146
  have disj: "?A' \<inter> ?B' = {}" by blast
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1147
  have "card ?C = card(?A' \<union> ?B')"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1148
    using uni by simp
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1149
  also have "\<dots> = card ?A + card ?B"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1150
    using card_Un_disjoint[OF _ _ disj] bij_betw_finite[OF f] bij_betw_finite[OF g]
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1151
      bij_betw_same_card[OF f] bij_betw_same_card[OF g] fin_A fin_B
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1152
    by presburger
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1153
  finally show ?thesis .
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1154
qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1155
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1156
lemma card_length_sum_list: "card {l::nat list. size l = m \<and> sum_list l = N} = (N + m - 1) choose N"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1157
  \<comment> \<open>by Holden Lee, tidied by Tobias Nipkow\<close>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1158
proof (cases m)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1159
  case 0
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1160
  then show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1161
    by (cases N) (auto cong: conj_cong)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1162
next
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1163
  case (Suc m')
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1164
  have m: "m \<ge> 1"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1165
    by (simp add: Suc)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1166
  then show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1167
  proof (induct "N + m - 1" arbitrary: N m)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1168
    case 0  \<comment> \<open>In the base case, the only solution is [0].\<close>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1169
    have [simp]: "{l::nat list. length l = Suc 0 \<and> (\<forall>n\<in>set l. n = 0)} = {[0]}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1170
      by (auto simp: length_Suc_conv)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1171
    have "m = 1 \<and> N = 0"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1172
      using 0 by linarith
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1173
    then show ?case
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1174
      by simp
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1175
  next
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1176
    case (Suc k)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1177
    have c1: "card {l::nat list. size l = (m - 1) \<and> sum_list l =  N} = (N + (m - 1) - 1) choose N"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1178
    proof (cases "m = 1")
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1179
      case True
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1180
      with Suc.hyps have "N \<ge> 1"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1181
        by auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1182
      with True show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1183
        by (simp add: binomial_eq_0)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1184
    next
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1185
      case False
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1186
      then show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1187
        using Suc by fastforce
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1188
    qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1189
    from Suc have c2: "card {l::nat list. size l = m \<and> sum_list l + 1 = N} =
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1190
      (if N > 0 then ((N - 1) + m - 1) choose (N - 1) else 0)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1191
    proof -
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1192
      have *: "n > 0 \<Longrightarrow> Suc m = n \<longleftrightarrow> m = n - 1" for m n
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1193
        by arith
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1194
      from Suc have "N > 0 \<Longrightarrow>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1195
        card {l::nat list. size l = m \<and> sum_list l + 1 = N} =
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1196
          ((N - 1) + m - 1) choose (N - 1)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1197
        by (simp add: *)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1198
      then show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1199
        by auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1200
    qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1201
    from Suc.prems have "(card {l::nat list. size l = (m - 1) \<and> sum_list l = N} +
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1202
          card {l::nat list. size l = m \<and> sum_list l + 1 = N}) = (N + m - 1) choose N"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1203
      by (auto simp: c1 c2 choose_reduce_nat[of "N + m - 1" N] simp del: One_nat_def)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1204
    then show ?case
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1205
      using card_length_sum_list_rec[OF Suc.prems] by auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1206
  qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1207
qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1208
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1209
lemma card_disjoint_shuffles:
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1210
  assumes "set xs \<inter> set ys = {}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1211
  shows   "card (shuffles xs ys) = (length xs + length ys) choose length xs"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1212
using assms
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1213
proof (induction xs ys rule: shuffles.induct)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1214
  case (3 x xs y ys)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1215
  have "shuffles (x # xs) (y # ys) = (#) x ` shuffles xs (y # ys) \<union> (#) y ` shuffles (x # xs) ys"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1216
    by (rule shuffles.simps)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1217
  also have "card \<dots> = card ((#) x ` shuffles xs (y # ys)) + card ((#) y ` shuffles (x # xs) ys)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1218
    by (rule card_Un_disjoint) (insert "3.prems", auto)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1219
  also have "card ((#) x ` shuffles xs (y # ys)) = card (shuffles xs (y # ys))"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1220
    by (rule card_image) auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1221
  also have "\<dots> = (length xs + length (y # ys)) choose length xs"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1222
    using "3.prems" by (intro "3.IH") auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1223
  also have "card ((#) y ` shuffles (x # xs) ys) = card (shuffles (x # xs) ys)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1224
    by (rule card_image) auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1225
  also have "\<dots> = (length (x # xs) + length ys) choose length (x # xs)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1226
    using "3.prems" by (intro "3.IH") auto
80175
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 79586
diff changeset
  1227
  also have "(length xs + length (y # ys) choose length xs) + \<dots> =
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1228
               (length (x # xs) + length (y # ys)) choose length (x # xs)" by simp
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1229
  finally show ?case .
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1230
qed auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1231
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1232
lemma Suc_times_binomial_add: "Suc a * (Suc (a + b) choose Suc a) = Suc b * (Suc (a + b) choose a)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1233
  \<comment> \<open>by Lukas Bulwahn\<close>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1234
proof -
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1235
  have dvd: "Suc a * (fact a * fact b) dvd fact (Suc (a + b))" for a b
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1236
    using fact_fact_dvd_fact[of "Suc a" "b", where 'a=nat]
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1237
    by (simp only: fact_Suc add_Suc[symmetric] of_nat_id mult.assoc)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1238
  have "Suc a * (fact (Suc (a + b)) div (Suc a * fact a * fact b)) =
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1239
      Suc a * fact (Suc (a + b)) div (Suc a * (fact a * fact b))"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1240
    by (subst div_mult_swap[symmetric]; simp only: mult.assoc dvd)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1241
  also have "\<dots> = Suc b * fact (Suc (a + b)) div (Suc b * (fact a * fact b))"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1242
    by (simp only: div_mult_mult1)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1243
  also have "\<dots> = Suc b * (fact (Suc (a + b)) div (Suc b * (fact a * fact b)))"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1244
    using dvd[of b a] by (subst div_mult_swap[symmetric]; simp only: ac_simps dvd)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1245
  finally show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1246
    by (subst (1 2) binomial_altdef_nat)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1247
      (simp_all only: ac_simps diff_Suc_Suc Suc_diff_le diff_add_inverse fact_Suc of_nat_id)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1248
qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1249
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1250
subsection \<open>Inclusion-exclusion principle\<close>
80175
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 79586
diff changeset
  1251
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1252
text \<open>Ported from HOL Light by lcp\<close>
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1253
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1254
lemma Inter_over_Union:
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1255
  "\<Inter> {\<Union> (\<F> x) |x. x \<in> S} = \<Union> {\<Inter> (G ` S) |G. \<forall>x\<in>S. G x \<in> \<F> x}" 
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1256
proof -
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1257
  have "\<And>x. \<forall>s\<in>S. \<exists>X \<in> \<F> s. x \<in> X \<Longrightarrow> \<exists>G. (\<forall>x\<in>S. G x \<in> \<F> x) \<and> (\<forall>s\<in>S. x \<in> G s)"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1258
    by metis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1259
  then show ?thesis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1260
    by (auto simp flip: all_simps ex_simps)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1261
qed
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1262
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1263
lemma subset_insert_lemma:
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1264
  "{T. T \<subseteq> (insert a S) \<and> P T} = {T. T \<subseteq> S \<and> P T} \<union> {insert a T |T. T \<subseteq> S \<and> P(insert a T)}" (is "?L=?R")
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1265
proof
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1266
  show "?L \<subseteq> ?R"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1267
    by (smt (verit) UnI1 UnI2 insert_Diff mem_Collect_eq subsetI subset_insert_iff)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1268
qed blast
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1269
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1270
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1271
text\<open>Versions for additive real functions, where the additivity applies only to some
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1272
 specific subsets (e.g. cardinality of finite sets, measurable sets with bounded measure.
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1273
 (From HOL Light)\<close>
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1274
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1275
locale Incl_Excl =
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1276
  fixes P :: "'a set \<Rightarrow> bool" and f :: "'a set \<Rightarrow> 'b::ring_1"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1277
  assumes disj_add: "\<lbrakk>P S; P T; disjnt S T\<rbrakk> \<Longrightarrow> f(S \<union> T) = f S + f T"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1278
    and empty: "P{}"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1279
    and Int: "\<lbrakk>P S; P T\<rbrakk> \<Longrightarrow> P(S \<inter> T)"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1280
    and Un: "\<lbrakk>P S; P T\<rbrakk> \<Longrightarrow> P(S \<union> T)"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1281
    and Diff: "\<lbrakk>P S; P T\<rbrakk> \<Longrightarrow> P(S - T)"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1282
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1283
begin
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1284
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1285
lemma f_empty [simp]: "f{} = 0"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1286
  using disj_add empty by fastforce
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1287
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1288
lemma f_Un_Int: "\<lbrakk>P S; P T\<rbrakk> \<Longrightarrow> f(S \<union> T) + f(S \<inter> T) = f S + f T"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1289
  by (smt (verit, ccfv_threshold) Groups.add_ac(2) Incl_Excl.Diff Incl_Excl.Int Incl_Excl_axioms Int_Diff_Un Int_Diff_disjoint Int_absorb Un_Diff Un_Int_eq(2) disj_add disjnt_def group_cancel.add2 sup_bot.right_neutral)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1290
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1291
lemma restricted_indexed:
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1292
  assumes "finite A" and X: "\<And>a. a \<in> A \<Longrightarrow> P(X a)"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1293
  shows "f(\<Union>(X ` A)) = (\<Sum>B | B \<subseteq> A \<and> B \<noteq> {}. (- 1) ^ (card B + 1) * f (\<Inter> (X ` B)))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1294
proof -
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1295
  have "\<lbrakk>finite A; card A = n; \<forall>a \<in> A. P (X a)\<rbrakk>
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1296
              \<Longrightarrow> f(\<Union>(X ` A)) = (\<Sum>B | B \<subseteq> A \<and> B \<noteq> {}. (- 1) ^ (card B + 1) * f (\<Inter> (X ` B)))" for n X and A :: "'c set"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1297
  proof (induction n arbitrary: A X rule: less_induct)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1298
    case (less n0 A0 X)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1299
    show ?case
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1300
    proof (cases "n0=0")
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1301
      case True
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1302
      with less show ?thesis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1303
       by fastforce
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1304
    next
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1305
      case False
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1306
      with less.prems obtain A n a where *: "n0 = Suc n" "A0 = insert a A" "a \<notin> A" "card A = n" "finite A"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1307
        by (metis card_Suc_eq_finite not0_implies_Suc)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1308
      with less have "P (X a)" by blast
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1309
      have APX: "\<forall>a \<in> A. P (X a)"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1310
        by (simp add: "*" less.prems)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1311
      have PUXA: "P (\<Union> (X ` A))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1312
        using \<open>finite A\<close> APX
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1313
        by (induction) (auto simp: empty Un)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1314
      have "f (\<Union> (X ` A0)) = f (X a \<union> \<Union> (X ` A))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1315
        by (simp add: *)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1316
      also have "... = f (X a) + f (\<Union> (X ` A)) - f (X a \<inter> \<Union> (X ` A))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1317
        using f_Un_Int add_diff_cancel PUXA \<open>P (X a)\<close> by metis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1318
      also have "... = f (X a) - (\<Sum>B | B \<subseteq> A \<and> B \<noteq> {}. (- 1) ^ card B * f (\<Inter> (X ` B))) +
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1319
                       (\<Sum>B | B \<subseteq> A \<and> B \<noteq> {}. (- 1) ^ card B * f (X a \<inter> \<Inter> (X ` B)))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1320
      proof -
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1321
        have 1: "f (\<Union>i\<in>A. X a \<inter> X i) = (\<Sum>B | B \<subseteq> A \<and> B \<noteq> {}. (- 1) ^ (card B + 1) * f (\<Inter>b\<in>B. X a \<inter> X b))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1322
          using less.IH [of n A "\<lambda>i. X a \<inter> X i"] APX Int \<open>P (X a)\<close>  by (simp add: *)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1323
        have 2: "X a \<inter> \<Union> (X ` A) = (\<Union>i\<in>A. X a \<inter> X i)"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1324
          by auto
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1325
        have 3: "f (\<Union> (X ` A)) = (\<Sum>B | B \<subseteq> A \<and> B \<noteq> {}. (- 1) ^ (card B + 1) * f (\<Inter> (X ` B)))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1326
          using less.IH [of n A X] APX Int \<open>P (X a)\<close>  by (simp add: *)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1327
        show ?thesis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1328
          unfolding 3 2 1
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1329
          by (simp add: sum_negf)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1330
      qed
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1331
      also have "... = (\<Sum>B | B \<subseteq> A0 \<and> B \<noteq> {}. (- 1) ^ (card B + 1) * f (\<Inter> (X ` B)))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1332
      proof -
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1333
         have F: "{insert a B |B. B \<subseteq> A} = insert a ` Pow A \<and> {B. B \<subseteq> A \<and> B \<noteq> {}} = Pow A - {{}}"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1334
          by auto
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1335
        have G: "(\<Sum>B\<in>Pow A. (- 1) ^ card (insert a B) * f (X a \<inter> \<Inter> (X ` B))) = (\<Sum>B\<in>Pow A. - ((- 1) ^ card B * f (X a \<inter> \<Inter> (X ` B))))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1336
        proof (rule sum.cong [OF refl])
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1337
          fix B
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1338
          assume B: "B \<in> Pow A"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1339
          then have "finite B"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1340
            using \<open>finite A\<close> finite_subset by auto
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1341
          show "(- 1) ^ card (insert a B) * f (X a \<inter> \<Inter> (X ` B)) = - ((- 1) ^ card B * f (X a \<inter> \<Inter> (X ` B)))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1342
            using B * by (auto simp add: card_insert_if \<open>finite B\<close>)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1343
        qed
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1344
        have disj: "{B. B \<subseteq> A \<and> B \<noteq> {}} \<inter> {insert a B |B. B \<subseteq> A} = {}"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1345
          using * by blast
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1346
        have inj: "inj_on (insert a) (Pow A)"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1347
          using "*" inj_on_def by fastforce
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1348
        show ?thesis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1349
          apply (simp add: * subset_insert_lemma sum.union_disjoint disj sum_negf)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1350
          apply (simp add: F G sum_negf sum.reindex [OF inj] o_def sum_diff *)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1351
          done
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1352
      qed
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1353
      finally show ?thesis .
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1354
    qed   
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1355
  qed
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1356
  then show ?thesis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1357
    by (meson assms)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1358
qed
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1359
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1360
lemma restricted:
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1361
  assumes "finite A" "\<And>a. a \<in> A \<Longrightarrow> P a"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1362
  shows  "f(\<Union> A) = (\<Sum>B | B \<subseteq> A \<and> B \<noteq> {}. (- 1) ^ (card B + 1) * f (\<Inter> B))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1363
  using restricted_indexed [of A "\<lambda>x. x"] assms by auto
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1364
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1365
end
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1366
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1367
subsection\<open>Versions for unrestrictedly additive functions\<close>
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1368
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1369
lemma Incl_Excl_UN:
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1370
  fixes f :: "'a set \<Rightarrow> 'b::ring_1"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1371
  assumes "\<And>S T. disjnt S T \<Longrightarrow> f(S \<union> T) = f S + f T" "finite A"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1372
  shows "f(\<Union>(G ` A)) = (\<Sum>B | B \<subseteq> A \<and> B \<noteq> {}. (-1) ^ (card B + 1) * f (\<Inter> (G ` B)))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1373
proof -
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1374
  interpret Incl_Excl "\<lambda>x. True" f
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1375
    by (simp add: Incl_Excl.intro assms(1))
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1376
  show ?thesis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1377
    using restricted_indexed assms by blast
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1378
qed
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1379
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1380
lemma Incl_Excl_Union:
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1381
  fixes f :: "'a set \<Rightarrow> 'b::ring_1"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1382
  assumes "\<And>S T. disjnt S T \<Longrightarrow> f(S \<union> T) = f S + f T" "finite A"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1383
  shows "f(\<Union> A) = (\<Sum>B | B \<subseteq> A \<and> B \<noteq> {}. (- 1) ^ (card B + 1) * f (\<Inter> B))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1384
  using Incl_Excl_UN[of f A "\<lambda>X. X"] assms by simp
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1385
75856
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1386
text \<open>The famous inclusion-exclusion formula for the cardinality of a union\<close>
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1387
lemma int_card_UNION:
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1388
  assumes "finite A" "\<And>K. K \<in> A \<Longrightarrow> finite K"
75856
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1389
  shows "int (card (\<Union>A)) = (\<Sum>I | I \<subseteq> A \<and> I \<noteq> {}. (- 1) ^ (card I + 1) * int (card (\<Inter>I)))"
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1390
proof -
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1391
  interpret Incl_Excl finite "int o card"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1392
  proof qed (auto simp add: card_Un_disjnt)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1393
  show ?thesis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1394
    using restricted assms by auto
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1395
qed
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1396
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1397
text\<open>A more conventional form\<close>
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1398
lemma inclusion_exclusion:
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1399
  assumes "finite A" "\<And>K. K \<in> A \<Longrightarrow> finite K"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1400
  shows "int(card(\<Union> A)) =
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1401
         (\<Sum>n=1..card A. (-1) ^ (Suc n) * (\<Sum>B | B \<subseteq> A \<and> card B = n. int (card (\<Inter> B))))" (is "_=?R")
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1402
proof -
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1403
  have fin: "finite {I. I \<subseteq> A \<and> I \<noteq> {}}"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1404
    by (simp add: assms)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1405
  have "\<And>k. \<lbrakk>Suc 0 \<le> k; k \<le> card A\<rbrakk> \<Longrightarrow> \<exists>B\<subseteq>A. B \<noteq> {} \<and> k = card B"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1406
    by (metis (mono_tags, lifting) Suc_le_D Zero_neq_Suc card_eq_0_iff obtain_subset_with_card_n)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1407
  with \<open>finite A\<close> finite_subset
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1408
  have card_eq: "card ` {I. I \<subseteq> A \<and> I \<noteq> {}} = {1..card A}"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1409
    using not_less_eq_eq card_mono by (fastforce simp: image_iff)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1410
  have "int(card(\<Union> A)) 
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1411
      = (\<Sum>y = 1..card A. \<Sum>I\<in>{x. x \<subseteq> A \<and> x \<noteq> {} \<and> card x = y}. - ((- 1) ^ y * int (card (\<Inter> I))))"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1412
    by (simp add: int_card_UNION assms sum.image_gen [OF fin, where g=card] card_eq)
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1413
  also have "... = ?R"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1414
  proof -
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1415
    have "{B. B \<subseteq> A \<and> B \<noteq> {} \<and> card B = k} = {B. B \<subseteq> A \<and> card B = k}"
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1416
      if "Suc 0 \<le> k" and "k \<le> card A" for k
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1417
      using that by auto
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1418
    then show ?thesis
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1419
      by (clarsimp simp add: sum_negf simp flip: sum_distrib_left)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1420
  qed
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1421
  finally show ?thesis .
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1422
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1423
75856
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1424
lemma card_UNION:
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1425
  assumes "finite A" and "\<And>K. K \<in> A \<Longrightarrow> finite K"
75856
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1426
  shows "card (\<Union>A) = nat (\<Sum>I | I \<subseteq> A \<and> I \<noteq> {}. (- 1) ^ (card I + 1) * int (card (\<Inter>I)))"
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1427
  by (simp only: flip: int_card_UNION [OF assms])
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1428
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1429
lemma card_UNION_nonneg:
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1430
  assumes "finite A" and "\<And>K. K \<in> A \<Longrightarrow> finite K"
75856
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1431
  shows "(\<Sum>I | I \<subseteq> A \<and> I \<noteq> {}. (- 1) ^ (card I + 1) * int (card (\<Inter>I))) \<ge> 0"
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1432
  using int_card_UNION [OF assms] by presburger
4b507edcf6b5 The right way to formulate card_UNION, plus the old version for compatibility
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  1433
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1434
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1435
subsection \<open>General "Moebius inversion" inclusion-exclusion principle\<close>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1436
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1437
text \<open>This "symmetric" form is from Ira Gessel: "Symmetric Inclusion-Exclusion" \<close>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1438
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1439
lemma sum_Un_eq:
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1440
   "\<lbrakk>S \<inter> T = {}; S \<union> T = U; finite U\<rbrakk>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1441
           \<Longrightarrow> (sum f S + sum f T = sum f U)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1442
  by (metis finite_Un sum.union_disjoint)
78656
4da1e18a9633 Loads of new material related to porting the Euler Polyhedron Formula from HOL Light
paulson <lp15@cam.ac.uk>
parents: 77172
diff changeset
  1443
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1444
lemma card_adjust_lemma: "\<lbrakk>inj_on f S; x = y + card (f ` S)\<rbrakk> \<Longrightarrow> x = y + card S"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1445
  by (simp add: card_image)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1446
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1447
lemma card_subsets_step:
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1448
  assumes "finite S" "x \<notin> S" "U \<subseteq> S"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1449
  shows "card {T. T \<subseteq> (insert x S) \<and> U \<subseteq> T \<and> odd(card T)} 
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1450
       = card {T. T \<subseteq> S \<and> U \<subseteq> T \<and> odd(card T)} + card {T. T \<subseteq> S \<and> U \<subseteq> T \<and> even(card T)} \<and>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1451
         card {T. T \<subseteq> (insert x S) \<and> U \<subseteq> T \<and> even(card T)} 
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1452
       = card {T. T \<subseteq> S \<and> U \<subseteq> T \<and> even(card T)} + card {T. T \<subseteq> S \<and> U \<subseteq> T \<and> odd(card T)}"
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1453
proof -
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1454
  have inj: "inj_on (insert x) {T. T \<subseteq> S \<and> P T}" for P
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1455
    using assms by (auto simp: inj_on_def)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1456
  have [simp]: "finite {T. T \<subseteq> S \<and> P T}"  "finite (insert x ` {T. T \<subseteq> S \<and> P T})" for P
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1457
    using \<open>finite S\<close> by auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1458
  have [simp]: "disjnt {T. T \<subseteq> S \<and> P T} (insert x ` {T. T \<subseteq> S \<and> Q T})" for P Q
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1459
    using assms by (auto simp: disjnt_iff)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1460
  have eq: "{T. T \<subseteq> S \<and> U \<subseteq> T \<and> P T} \<union> insert x ` {T. T \<subseteq> S \<and> U \<subseteq> T \<and> Q T}
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1461
          = {T. T \<subseteq> insert x S \<and> U \<subseteq> T \<and> P T}"  (is "?L = ?R")
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1462
    if "\<And>A. A \<subseteq> S \<Longrightarrow> Q (insert x A) \<longleftrightarrow> P A" "\<And>A. \<not> Q A \<longleftrightarrow> P A" for P Q
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1463
  proof
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1464
    show "?L \<subseteq> ?R"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1465
      by (clarsimp simp: image_iff subset_iff) (meson subsetI that)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1466
    show "?R \<subseteq> ?L"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1467
      using \<open>U \<subseteq> S\<close>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1468
      by (clarsimp simp: image_iff) (smt (verit) insert_iff mk_disjoint_insert subset_iff that)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1469
  qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1470
  have [simp]: "\<And>A. A \<subseteq> S \<Longrightarrow> even (card (insert x A)) \<longleftrightarrow> odd (card A)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1471
    by (metis \<open>finite S\<close> \<open>x \<notin> S\<close> card_insert_disjoint even_Suc finite_subset subsetD)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1472
  show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1473
    by (intro conjI card_adjust_lemma [OF inj]; simp add: eq flip: card_Un_disjnt)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1474
qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1475
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1476
lemma card_subsupersets_even_odd:
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1477
  assumes "finite S" "U \<subset> S"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1478
  shows "card {T. T \<subseteq> S \<and> U \<subseteq> T \<and> even(card T)} 
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1479
       = card {T. T \<subseteq> S \<and> U \<subseteq> T \<and> odd(card T)}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1480
  using assms
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1481
proof (induction "card S" arbitrary: S rule: less_induct)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1482
  case (less S)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1483
  then obtain x where "x \<notin> U" "x \<in> S"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1484
    by blast
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1485
  then have U: "U \<subseteq> S - {x}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1486
    using less.prems(2) by blast
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1487
  let ?V = "S - {x}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1488
  show ?case
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1489
    using card_subsets_step [of ?V x U] less.prems U
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1490
    by (simp add: insert_absorb \<open>x \<in> S\<close>)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1491
qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1492
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1493
lemma sum_alternating_cancels:
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1494
  assumes "finite S" "card {x. x \<in> S \<and> even(f x)} = card {x. x \<in> S \<and> odd(f x)}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1495
  shows "(\<Sum>x\<in>S. (-1) ^ f x) = (0::'b::ring_1)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1496
proof -
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1497
  have "(\<Sum>x\<in>S. (-1) ^ f x)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1498
      = (\<Sum>x | x \<in> S \<and> even (f x). (-1) ^ f x) + (\<Sum>x | x \<in> S \<and> odd (f x). (-1) ^ f x)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1499
    by (rule sum_Un_eq [symmetric]; force simp: \<open>finite S\<close>)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1500
  also have "... = (0::'b::ring_1)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1501
    by (simp add: minus_one_power_iff assms cong: conj_cong)
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1502
  finally show ?thesis .
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1503
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1504
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1505
lemma inclusion_exclusion_symmetric:
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1506
  fixes f :: "'a set \<Rightarrow> 'b::ring_1"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1507
  assumes \<section>: "\<And>S. finite S \<Longrightarrow> g S = (\<Sum>T \<in> Pow S. (-1) ^ card T * f T)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1508
    and "finite S"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1509
  shows "f S = (\<Sum>T \<in> Pow S. (-1) ^ card T * g T)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1510
proof -
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1511
  have "(-1) ^ card T * g T = (-1) ^ card T * (\<Sum>U | U \<subseteq> S \<and> U \<subseteq> T. (-1) ^ card U * f U)" 
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1512
    if "T \<subseteq> S" for T
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1513
  proof -
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1514
    have [simp]: "{U. U \<subseteq> S \<and> U \<subseteq> T} = Pow T"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1515
      using that by auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1516
    show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1517
      using that by (simp add: \<open>finite S\<close> finite_subset \<section>)
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1518
  qed
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1519
  then have "(\<Sum>T \<in> Pow S. (-1) ^ card T * g T)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1520
      = (\<Sum>T\<in>Pow S. (-1) ^ card T * (\<Sum>U | U \<in> {U. U \<subseteq> S} \<and> U \<subseteq> T. (-1) ^ card U * f U))"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1521
    by simp
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1522
  also have "... = (\<Sum>U\<in>Pow S. (\<Sum>T | T \<subseteq> S \<and> U \<subseteq> T. (-1) ^ card T) * (-1) ^ card U * f U)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1523
    unfolding sum_distrib_left
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1524
    by (subst sum.swap_restrict; simp add: \<open>finite S\<close> algebra_simps sum_distrib_right Pow_def)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1525
  also have "... = (\<Sum>U\<in>Pow S. if U=S then f S else 0)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1526
  proof -
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1527
    have [simp]: "{T. T \<subseteq> S \<and> S \<subseteq> T} = {S}"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1528
      by auto
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1529
    show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1530
      apply (rule sum.cong [OF refl])
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1531
      by (simp add: sum_alternating_cancels card_subsupersets_even_odd \<open>finite S\<close> flip: power_add)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1532
  qed
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1533
  also have "... = f S"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1534
    by (simp add: \<open>finite S\<close>)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1535
  finally show ?thesis
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1536
    by presburger
59667
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1537
qed
651ea265d568 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents: 59658
diff changeset
  1538
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1539
text\<open> The more typical non-symmetric version.                                   \<close>
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1540
lemma inclusion_exclusion_mobius:
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1541
  fixes f :: "'a set \<Rightarrow> 'b::ring_1"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1542
  assumes \<section>: "\<And>S. finite S \<Longrightarrow> g S = sum f (Pow S)" and "finite S"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1543
  shows "f S = (\<Sum>T \<in> Pow S. (-1) ^ (card S - card T) * g T)" (is "_ = ?rhs")
60604
dd4253d5dd82 tuned src/HOL/ex/Ballot
hoelzl
parents: 60301
diff changeset
  1544
proof -
78667
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1545
  have "(- 1) ^ card S * f S = (\<Sum>T\<in>Pow S. (- 1) ^ card T * g T)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1546
    by (rule inclusion_exclusion_symmetric; simp add: assms flip: power_add mult.assoc)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1547
  then have "((- 1) ^ card S * (- 1) ^ card S) * f S = ((- 1) ^ card S) * (\<Sum>T\<in>Pow S. (- 1) ^ card T * g T)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1548
    by (simp add: mult_ac)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1549
  then have "f S = (\<Sum>T\<in>Pow S. (- 1) ^ (card S + card T) * g T)"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1550
    by (simp add: sum_distrib_left flip: power_add mult.assoc)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1551
  also have "... = ?rhs"
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1552
    by (simp add: \<open>finite S\<close> card_mono neg_one_power_add_eq_neg_one_power_diff)
d900ff3f314a A few more inclusion-exclusion theorems from HOL Light
paulson <lp15@cam.ac.uk>
parents: 78656
diff changeset
  1553
  finally show ?thesis .
60604
dd4253d5dd82 tuned src/HOL/ex/Ballot
hoelzl
parents: 60301
diff changeset
  1554
qed
dd4253d5dd82 tuned src/HOL/ex/Ballot
hoelzl
parents: 60301
diff changeset
  1555
63373
487d764fca4a tuned sections
haftmann
parents: 63372
diff changeset
  1556
68785
862b1288020f tuned code setup
haftmann
parents: 68784
diff changeset
  1557
subsection \<open>Executable code\<close>
63373
487d764fca4a tuned sections
haftmann
parents: 63372
diff changeset
  1558
62128
3201ddb00097 Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents: 61799
diff changeset
  1559
lemma gbinomial_code [code]:
68787
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
  1560
  "a gchoose k =
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
  1561
    (if k = 0 then 1
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
  1562
     else fold_atLeastAtMost_nat (\<lambda>k acc. (a - of_nat k) * acc) 0 (k - 1) 1 / fact k)"
b129052644e9 more uniform parameter naming convention for choose and gchoose
haftmann
parents: 68786
diff changeset
  1563
  by (cases k)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1564
    (simp_all add: gbinomial_prod_rev prod_atLeastAtMost_code [symmetric]
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1565
      atLeastLessThanSuc_atLeastAtMost)
62128
3201ddb00097 Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents: 61799
diff changeset
  1566
3201ddb00097 Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents: 61799
diff changeset
  1567
lemma binomial_code [code]:
68785
862b1288020f tuned code setup
haftmann
parents: 68784
diff changeset
  1568
  "n choose k =
62128
3201ddb00097 Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents: 61799
diff changeset
  1569
      (if k > n then 0
68785
862b1288020f tuned code setup
haftmann
parents: 68784
diff changeset
  1570
       else if 2 * k > n then n choose (n - k)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68787
diff changeset
  1571
       else (fold_atLeastAtMost_nat (*) (n - k + 1) n 1 div fact k))"
62128
3201ddb00097 Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents: 61799
diff changeset
  1572
proof -
3201ddb00097 Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents: 61799
diff changeset
  1573
  {
3201ddb00097 Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents: 61799
diff changeset
  1574
    assume "k \<le> n"
63466
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1575
    then have "{1..n} = {1..n-k} \<union> {n-k+1..n}" by auto
2100fbbdc3f1 misc tuning and modernization;
wenzelm
parents: 63417
diff changeset
  1576
    then have "(fact n :: nat) = fact (n-k) * \<Prod>{n-k+1..n}"
65581
baf96277ee76 better code equation for binomial
eberlm <eberlm@in.tum.de>
parents: 65552
diff changeset
  1577
      by (simp add: prod.union_disjoint fact_prod)
62128
3201ddb00097 Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents: 61799
diff changeset
  1578
  }
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1579
  then show ?thesis by (auto simp: binomial_altdef_nat mult_ac prod_atLeastAtMost_code)
62378
85ed00c1fe7c generalize more theorems to support enat and ennreal
hoelzl
parents: 62347
diff changeset
  1580
qed
62128
3201ddb00097 Integrated some material from Algebraic_Numbers AFP entry to Polynomials; generalised some polynomial stuff.
eberlm
parents: 61799
diff changeset
  1581
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15094
diff changeset
  1582
end