| author | nipkow | 
| Tue, 11 Oct 2022 10:45:42 +0200 | |
| changeset 76259 | d1c26efb7a47 | 
| parent 75865 | 62c64e3e4741 | 
| child 79566 | f783490c6c99 | 
| permissions | -rw-r--r-- | 
| 66589 | 1 | (* Title: HOL/Factorial.thy | 
| 65812 | 2 | Author: Jacques D. Fleuriot | 
| 3 | Author: Lawrence C Paulson | |
| 4 | Author: Jeremy Avigad | |
| 5 | Author: Chaitanya Mangla | |
| 6 | Author: Manuel Eberl | |
| 7 | *) | |
| 8 | ||
| 9 | section \<open>Factorial Function, Rising Factorials\<close> | |
| 10 | ||
| 11 | theory Factorial | |
| 65813 | 12 | imports Groups_List | 
| 65812 | 13 | begin | 
| 14 | ||
| 15 | subsection \<open>Factorial Function\<close> | |
| 16 | ||
| 17 | context semiring_char_0 | |
| 18 | begin | |
| 19 | ||
| 20 | definition fact :: "nat \<Rightarrow> 'a" | |
| 21 |   where fact_prod: "fact n = of_nat (\<Prod>{1..n})"
 | |
| 22 | ||
| 23 | lemma fact_prod_Suc: "fact n = of_nat (prod Suc {0..<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: 
69182diff
changeset | 24 | unfolding fact_prod using atLeast0LessThan prod.atLeast1_atMost_eq by auto | 
| 65812 | 25 | |
| 26 | lemma fact_prod_rev: "fact n = of_nat (\<Prod>i = 0..<n. n - 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: 
69182diff
changeset | 27 | proof - | 
| 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69182diff
changeset | 28 |   have "prod Suc {0..<n} = \<Prod>{1..n}"
 | 
| 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69182diff
changeset | 29 | by (simp add: atLeast0LessThan prod.atLeast1_atMost_eq) | 
| 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69182diff
changeset | 30 |   then have "prod Suc {0..<n} = prod ((-) (n + 1)) {1..n}"
 | 
| 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69182diff
changeset | 31 | using prod.atLeastAtMost_rev [of "\<lambda>i. i" 1 n] by presburger | 
| 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69182diff
changeset | 32 | then show ?thesis | 
| 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69182diff
changeset | 33 | unfolding fact_prod_Suc by (simp add: atLeast0LessThan prod.atLeast1_atMost_eq) | 
| 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69182diff
changeset | 34 | qed | 
| 65812 | 35 | |
| 36 | lemma fact_0 [simp]: "fact 0 = 1" | |
| 37 | by (simp add: fact_prod) | |
| 38 | ||
| 39 | lemma fact_1 [simp]: "fact 1 = 1" | |
| 40 | by (simp add: fact_prod) | |
| 41 | ||
| 42 | lemma fact_Suc_0 [simp]: "fact (Suc 0) = 1" | |
| 43 | by (simp add: fact_prod) | |
| 44 | ||
| 45 | lemma fact_Suc [simp]: "fact (Suc n) = of_nat (Suc n) * fact n" | |
| 46 | by (simp add: fact_prod atLeastAtMostSuc_conv algebra_simps) | |
| 47 | ||
| 48 | lemma fact_2 [simp]: "fact 2 = 2" | |
| 49 | by (simp add: numeral_2_eq_2) | |
| 50 | ||
| 51 | lemma fact_split: "k \<le> n \<Longrightarrow> fact n = of_nat (prod Suc {n - k..<n}) * fact (n - k)"
 | |
| 52 | by (simp add: fact_prod_Suc prod.union_disjoint [symmetric] | |
| 53 | ivl_disj_un ac_simps of_nat_mult [symmetric]) | |
| 54 | ||
| 55 | end | |
| 56 | ||
| 57 | lemma of_nat_fact [simp]: "of_nat (fact n) = fact n" | |
| 58 | by (simp add: fact_prod) | |
| 59 | ||
| 60 | lemma of_int_fact [simp]: "of_int (fact n) = fact n" | |
| 61 | by (simp only: fact_prod of_int_of_nat_eq) | |
| 62 | ||
| 63 | lemma fact_reduce: "n > 0 \<Longrightarrow> fact n = of_nat n * fact (n - 1)" | |
| 64 | by (cases n) auto | |
| 65 | ||
| 66 | lemma fact_nonzero [simp]: "fact n \<noteq> (0::'a::{semiring_char_0,semiring_no_zero_divisors})"
 | |
| 75865 
62c64e3e4741
The same, without adding a new simprule
 paulson <lp15@cam.ac.uk> parents: 
74969diff
changeset | 67 | using of_nat_0_neq by (induct n) auto | 
| 65812 | 68 | |
| 69 | lemma fact_mono_nat: "m \<le> n \<Longrightarrow> fact m \<le> (fact n :: nat)" | |
| 70 | by (induct n) (auto simp: le_Suc_eq) | |
| 71 | ||
| 72 | lemma fact_in_Nats: "fact n \<in> \<nat>" | |
| 73 | by (induct n) auto | |
| 74 | ||
| 75 | lemma fact_in_Ints: "fact n \<in> \<int>" | |
| 76 | by (induct n) auto | |
| 77 | ||
| 78 | context | |
| 79 |   assumes "SORT_CONSTRAINT('a::linordered_semidom)"
 | |
| 80 | begin | |
| 81 | ||
| 82 | lemma fact_mono: "m \<le> n \<Longrightarrow> fact m \<le> (fact n :: 'a)" | |
| 83 | by (metis of_nat_fact of_nat_le_iff fact_mono_nat) | |
| 84 | ||
| 85 | lemma fact_ge_1 [simp]: "fact n \<ge> (1 :: 'a)" | |
| 86 | by (metis le0 fact_0 fact_mono) | |
| 87 | ||
| 88 | lemma fact_gt_zero [simp]: "fact n > (0 :: 'a)" | |
| 89 | using fact_ge_1 less_le_trans zero_less_one by blast | |
| 90 | ||
| 91 | lemma fact_ge_zero [simp]: "fact n \<ge> (0 :: 'a)" | |
| 92 | by (simp add: less_imp_le) | |
| 93 | ||
| 94 | lemma fact_not_neg [simp]: "\<not> fact n < (0 :: 'a)" | |
| 95 | by (simp add: not_less_iff_gr_or_eq) | |
| 96 | ||
| 97 | lemma fact_le_power: "fact n \<le> (of_nat (n^n) :: 'a)" | |
| 98 | proof (induct n) | |
| 99 | case 0 | |
| 100 | then show ?case by simp | |
| 101 | next | |
| 102 | case (Suc n) | |
| 103 | then have *: "fact n \<le> (of_nat (Suc n ^ n) ::'a)" | |
| 104 | by (rule order_trans) (simp add: power_mono del: of_nat_power) | |
| 105 | have "fact (Suc n) = (of_nat (Suc n) * fact n ::'a)" | |
| 106 | by (simp add: algebra_simps) | |
| 107 | also have "\<dots> \<le> of_nat (Suc n) * of_nat (Suc n ^ n)" | |
| 108 | by (simp add: * ordered_comm_semiring_class.comm_mult_left_mono del: of_nat_power) | |
| 109 | also have "\<dots> \<le> of_nat (Suc n ^ Suc n)" | |
| 110 | by (metis of_nat_mult order_refl power_Suc) | |
| 111 | finally show ?case . | |
| 112 | qed | |
| 113 | ||
| 114 | end | |
| 115 | ||
| 116 | lemma fact_less_mono_nat: "0 < m \<Longrightarrow> m < n \<Longrightarrow> fact m < (fact n :: nat)" | |
| 117 | by (induct n) (auto simp: less_Suc_eq) | |
| 118 | ||
| 119 | lemma fact_less_mono: "0 < m \<Longrightarrow> m < n \<Longrightarrow> fact m < (fact n :: 'a::linordered_semidom)" | |
| 120 | by (metis of_nat_fact of_nat_less_iff fact_less_mono_nat) | |
| 121 | ||
| 122 | lemma fact_ge_Suc_0_nat [simp]: "fact n \<ge> Suc 0" | |
| 123 | by (metis One_nat_def fact_ge_1) | |
| 124 | ||
| 125 | lemma dvd_fact: "1 \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> m dvd fact n" | |
| 126 | by (induct n) (auto simp: dvdI le_Suc_eq) | |
| 127 | ||
| 128 | lemma fact_ge_self: "fact n \<ge> n" | |
| 129 | by (cases "n = 0") (simp_all add: dvd_imp_le dvd_fact) | |
| 130 | ||
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66589diff
changeset | 131 | lemma fact_dvd: "n \<le> m \<Longrightarrow> fact n dvd (fact m :: 'a::linordered_semidom)" | 
| 65812 | 132 | by (induct m) (auto simp: le_Suc_eq) | 
| 133 | ||
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66589diff
changeset | 134 | lemma fact_mod: "m \<le> n \<Longrightarrow> fact n mod (fact m :: 'a::{semidom_modulo, linordered_semidom}) = 0"
 | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66589diff
changeset | 135 | by (simp add: mod_eq_0_iff_dvd fact_dvd) | 
| 65812 | 136 | |
| 74969 | 137 | lemma fact_eq_fact_times: | 
| 138 | assumes "m \<ge> n" | |
| 139 |   shows "fact m = fact n * \<Prod>{Suc n..m}"
 | |
| 140 | unfolding fact_prod | |
| 141 | by (metis add.commute assms le_add1 le_add_diff_inverse of_nat_id plus_1_eq_Suc prod.ub_add_nat) | |
| 142 | ||
| 65812 | 143 | lemma fact_div_fact: | 
| 144 | assumes "m \<ge> n" | |
| 145 |   shows "fact m div fact n = \<Prod>{n + 1..m}"
 | |
| 74969 | 146 | by (simp add: fact_eq_fact_times [OF assms]) | 
| 65812 | 147 | |
| 148 | lemma fact_num_eq_if: "fact m = (if m = 0 then 1 else of_nat m * fact (m - 1))" | |
| 149 | by (cases m) auto | |
| 150 | ||
| 151 | lemma fact_div_fact_le_pow: | |
| 152 | assumes "r \<le> n" | |
| 153 | shows "fact n div fact (n - r) \<le> n ^ r" | |
| 154 | proof - | |
| 155 |   have "r \<le> n \<Longrightarrow> \<Prod>{n - r..n} = (n - r) * \<Prod>{Suc (n - r)..n}" for r
 | |
| 156 | by (subst prod.insert[symmetric]) (auto simp: atLeastAtMost_insertL) | |
| 157 | with assms show ?thesis | |
| 158 | by (induct r rule: nat.induct) (auto simp add: fact_div_fact Suc_diff_Suc mult_le_mono) | |
| 159 | qed | |
| 160 | ||
| 68783 | 161 | lemma prod_Suc_fact: "prod Suc {0..<n} = fact n"
 | 
| 162 | by (simp add: fact_prod_Suc) | |
| 163 | ||
| 164 | lemma prod_Suc_Suc_fact: "prod Suc {Suc 0..<n} = fact n"
 | |
| 165 | proof (cases "n = 0") | |
| 166 | case True | |
| 167 | then show ?thesis by simp | |
| 168 | next | |
| 169 | case False | |
| 170 |   have "prod Suc {Suc 0..<n} = Suc 0 * prod Suc {Suc 0..<n}"
 | |
| 171 | by simp | |
| 172 |   also have "\<dots> = prod Suc (insert 0 {Suc 0..<n})"
 | |
| 173 | by simp | |
| 174 |   also have "insert 0 {Suc 0..<n} = {0..<n}"
 | |
| 175 | using False by auto | |
| 176 | finally show ?thesis | |
| 177 | by (simp add: fact_prod_Suc) | |
| 178 | qed | |
| 179 | ||
| 65812 | 180 | lemma fact_numeral: "fact (numeral k) = numeral k * fact (pred_numeral k)" | 
| 181 | \<comment> \<open>Evaluation for specific numerals\<close> | |
| 182 | by (metis fact_Suc numeral_eq_Suc of_nat_numeral) | |
| 183 | ||
| 184 | ||
| 185 | subsection \<open>Pochhammer's symbol: generalized rising factorial\<close> | |
| 186 | ||
| 68224 | 187 | text \<open>See \<^url>\<open>https://en.wikipedia.org/wiki/Pochhammer_symbol\<close>.\<close> | 
| 65812 | 188 | |
| 189 | context comm_semiring_1 | |
| 190 | begin | |
| 191 | ||
| 192 | definition pochhammer :: "'a \<Rightarrow> nat \<Rightarrow> 'a" | |
| 193 |   where pochhammer_prod: "pochhammer a n = prod (\<lambda>i. a + of_nat i) {0..<n}"
 | |
| 194 | ||
| 195 | lemma pochhammer_prod_rev: "pochhammer a n = prod (\<lambda>i. a + of_nat (n - i)) {1..n}"
 | |
| 67411 
3f4b0c84630f
restored naming of lemmas after corresponding constants
 haftmann parents: 
67399diff
changeset | 196 | using prod.atLeastLessThan_rev_at_least_Suc_atMost [of "\<lambda>i. a + of_nat i" 0 n] | 
| 65812 | 197 | by (simp add: pochhammer_prod) | 
| 198 | ||
| 199 | lemma pochhammer_Suc_prod: "pochhammer a (Suc n) = prod (\<lambda>i. a + of_nat i) {0..n}"
 | |
| 200 | by (simp add: pochhammer_prod atLeastLessThanSuc_atLeastAtMost) | |
| 201 | ||
| 202 | lemma pochhammer_Suc_prod_rev: "pochhammer a (Suc n) = prod (\<lambda>i. a + of_nat (n - i)) {0..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: 
69182diff
changeset | 203 | using prod.atLeast_Suc_atMost_Suc_shift | 
| 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69182diff
changeset | 204 | by (simp add: pochhammer_prod_rev prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc) | 
| 65812 | 205 | |
| 206 | lemma pochhammer_0 [simp]: "pochhammer a 0 = 1" | |
| 207 | by (simp add: pochhammer_prod) | |
| 208 | ||
| 209 | lemma pochhammer_1 [simp]: "pochhammer a 1 = a" | |
| 210 | by (simp add: pochhammer_prod lessThan_Suc) | |
| 211 | ||
| 212 | lemma pochhammer_Suc0 [simp]: "pochhammer a (Suc 0) = a" | |
| 213 | by (simp add: pochhammer_prod lessThan_Suc) | |
| 214 | ||
| 215 | lemma pochhammer_Suc: "pochhammer a (Suc n) = pochhammer a n * (a + of_nat n)" | |
| 216 | by (simp add: pochhammer_prod atLeast0_lessThan_Suc ac_simps) | |
| 217 | ||
| 218 | end | |
| 219 | ||
| 220 | lemma pochhammer_nonneg: | |
| 221 | fixes x :: "'a :: linordered_semidom" | |
| 222 | shows "x > 0 \<Longrightarrow> pochhammer x n \<ge> 0" | |
| 223 | by (induction n) (auto simp: pochhammer_Suc intro!: mult_nonneg_nonneg add_nonneg_nonneg) | |
| 224 | ||
| 225 | lemma pochhammer_pos: | |
| 226 | fixes x :: "'a :: linordered_semidom" | |
| 227 | shows "x > 0 \<Longrightarrow> pochhammer x n > 0" | |
| 228 | by (induction n) (auto simp: pochhammer_Suc intro!: mult_pos_pos add_pos_nonneg) | |
| 229 | ||
| 69182 | 230 | context comm_semiring_1 | 
| 231 | begin | |
| 232 | ||
| 65812 | 233 | lemma pochhammer_of_nat: "pochhammer (of_nat x) n = of_nat (pochhammer x n)" | 
| 69182 | 234 | by (simp add: pochhammer_prod Factorial.pochhammer_prod) | 
| 235 | ||
| 236 | end | |
| 237 | ||
| 238 | context comm_ring_1 | |
| 239 | begin | |
| 65812 | 240 | |
| 241 | lemma pochhammer_of_int: "pochhammer (of_int x) n = of_int (pochhammer x n)" | |
| 69182 | 242 | by (simp add: pochhammer_prod Factorial.pochhammer_prod) | 
| 243 | ||
| 244 | end | |
| 65812 | 245 | |
| 246 | lemma pochhammer_rec: "pochhammer a (Suc n) = a * pochhammer (a + 1) 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: 
69182diff
changeset | 247 | by (simp add: pochhammer_prod prod.atLeast0_lessThan_Suc_shift ac_simps del: prod.op_ivl_Suc) | 
| 65812 | 248 | |
| 249 | lemma pochhammer_rec': "pochhammer z (Suc n) = (z + of_nat n) * pochhammer z n" | |
| 250 | by (simp add: pochhammer_prod prod.atLeast0_lessThan_Suc ac_simps) | |
| 251 | ||
| 252 | lemma pochhammer_fact: "fact n = pochhammer 1 n" | |
| 253 | by (simp add: pochhammer_prod fact_prod_Suc) | |
| 254 | ||
| 255 | lemma pochhammer_of_nat_eq_0_lemma: "k > n \<Longrightarrow> pochhammer (- (of_nat n :: 'a:: idom)) k = 0" | |
| 256 | by (auto simp add: pochhammer_prod) | |
| 257 | ||
| 258 | lemma pochhammer_of_nat_eq_0_lemma': | |
| 259 | assumes kn: "k \<le> n" | |
| 260 |   shows "pochhammer (- (of_nat n :: 'a::{idom,ring_char_0})) k \<noteq> 0"
 | |
| 261 | proof (cases k) | |
| 262 | case 0 | |
| 263 | then show ?thesis by simp | |
| 264 | next | |
| 265 | case (Suc h) | |
| 266 | then show ?thesis | |
| 267 | apply (simp add: pochhammer_Suc_prod) | |
| 268 | using Suc kn | |
| 269 | apply (auto simp add: algebra_simps) | |
| 270 | done | |
| 271 | qed | |
| 272 | ||
| 273 | lemma pochhammer_of_nat_eq_0_iff: | |
| 274 |   "pochhammer (- (of_nat n :: 'a::{idom,ring_char_0})) k = 0 \<longleftrightarrow> k > n"
 | |
| 275 | (is "?l = ?r") | |
| 276 | using pochhammer_of_nat_eq_0_lemma[of n k, where ?'a='a] | |
| 277 | pochhammer_of_nat_eq_0_lemma'[of k n, where ?'a = 'a] | |
| 278 | by (auto simp add: not_le[symmetric]) | |
| 279 | ||
| 66394 
32084d7e6b59
Some facts about orders of zeros
 eberlm <eberlm@in.tum.de> parents: 
65813diff
changeset | 280 | lemma pochhammer_0_left: | 
| 
32084d7e6b59
Some facts about orders of zeros
 eberlm <eberlm@in.tum.de> parents: 
65813diff
changeset | 281 | "pochhammer 0 n = (if n = 0 then 1 else 0)" | 
| 
32084d7e6b59
Some facts about orders of zeros
 eberlm <eberlm@in.tum.de> parents: 
65813diff
changeset | 282 | by (induction n) (simp_all add: pochhammer_rec) | 
| 
32084d7e6b59
Some facts about orders of zeros
 eberlm <eberlm@in.tum.de> parents: 
65813diff
changeset | 283 | |
| 65812 | 284 | lemma pochhammer_eq_0_iff: "pochhammer a n = (0::'a::field_char_0) \<longleftrightarrow> (\<exists>k < n. a = - of_nat k)" | 
| 285 | by (auto simp add: pochhammer_prod eq_neg_iff_add_eq_0) | |
| 286 | ||
| 287 | lemma pochhammer_eq_0_mono: | |
| 288 | "pochhammer a n = (0::'a::field_char_0) \<Longrightarrow> m \<ge> n \<Longrightarrow> pochhammer a m = 0" | |
| 289 | unfolding pochhammer_eq_0_iff by auto | |
| 290 | ||
| 291 | lemma pochhammer_neq_0_mono: | |
| 292 | "pochhammer a m \<noteq> (0::'a::field_char_0) \<Longrightarrow> m \<ge> n \<Longrightarrow> pochhammer a n \<noteq> 0" | |
| 293 | unfolding pochhammer_eq_0_iff by auto | |
| 294 | ||
| 295 | lemma pochhammer_minus: | |
| 296 | "pochhammer (- b) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (b - of_nat k + 1) k" | |
| 297 | proof (cases k) | |
| 298 | case 0 | |
| 299 | then show ?thesis by simp | |
| 300 | next | |
| 301 | case (Suc h) | |
| 302 | have eq: "((- 1) ^ Suc h :: 'a) = (\<Prod>i = 0..h. - 1)" | |
| 303 |     using prod_constant [where A="{0.. h}" and y="- 1 :: 'a"]
 | |
| 304 | by auto | |
| 305 | with Suc show ?thesis | |
| 67969 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 paulson <lp15@cam.ac.uk> parents: 
67411diff
changeset | 306 | using pochhammer_Suc_prod_rev [of "b - of_nat k + 1"] | 
| 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 paulson <lp15@cam.ac.uk> parents: 
67411diff
changeset | 307 | by (auto simp add: pochhammer_Suc_prod prod.distrib [symmetric] eq of_nat_diff simp del: prod_constant) | 
| 65812 | 308 | qed | 
| 309 | ||
| 310 | lemma pochhammer_minus': | |
| 311 | "pochhammer (b - of_nat k + 1) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (- b) k" | |
| 67969 
83c8cafdebe8
Syntax for the special cases Min(A`I) and Max (A`I)
 paulson <lp15@cam.ac.uk> parents: 
67411diff
changeset | 312 | by (simp add: pochhammer_minus) | 
| 65812 | 313 | |
| 314 | lemma pochhammer_same: "pochhammer (- of_nat n) n = | |
| 315 |     ((- 1) ^ n :: 'a::{semiring_char_0,comm_ring_1,semiring_no_zero_divisors}) * fact n"
 | |
| 316 | unfolding pochhammer_minus | |
| 317 | by (simp add: of_nat_diff pochhammer_fact) | |
| 318 | ||
| 319 | lemma pochhammer_product': "pochhammer z (n + m) = pochhammer z n * pochhammer (z + of_nat n) m" | |
| 320 | proof (induct n arbitrary: z) | |
| 321 | case 0 | |
| 322 | then show ?case by simp | |
| 323 | next | |
| 324 | case (Suc n z) | |
| 325 | have "pochhammer z (Suc n) * pochhammer (z + of_nat (Suc n)) m = | |
| 326 | z * (pochhammer (z + 1) n * pochhammer (z + 1 + of_nat n) m)" | |
| 327 | by (simp add: pochhammer_rec ac_simps) | |
| 328 | also note Suc[symmetric] | |
| 329 | also have "z * pochhammer (z + 1) (n + m) = pochhammer z (Suc (n + m))" | |
| 330 | by (subst pochhammer_rec) simp | |
| 331 | finally show ?case | |
| 332 | by simp | |
| 333 | qed | |
| 334 | ||
| 335 | lemma pochhammer_product: | |
| 336 | "m \<le> n \<Longrightarrow> pochhammer z n = pochhammer z m * pochhammer (z + of_nat m) (n - m)" | |
| 337 | using pochhammer_product'[of z m "n - m"] by simp | |
| 338 | ||
| 339 | lemma pochhammer_times_pochhammer_half: | |
| 340 | fixes z :: "'a::field_char_0" | |
| 341 | shows "pochhammer z (Suc n) * pochhammer (z + 1/2) (Suc n) = (\<Prod>k=0..2*n+1. z + of_nat k / 2)" | |
| 342 | proof (induct n) | |
| 343 | case 0 | |
| 344 | then show ?case | |
| 345 | by (simp add: atLeast0_atMost_Suc) | |
| 346 | next | |
| 347 | case (Suc n) | |
| 348 | define n' where "n' = Suc n" | |
| 349 | have "pochhammer z (Suc n') * pochhammer (z + 1 / 2) (Suc n') = | |
| 350 | (pochhammer z n' * pochhammer (z + 1 / 2) n') * ((z + of_nat n') * (z + 1/2 + of_nat n'))" | |
| 351 | (is "_ = _ * ?A") | |
| 352 | by (simp_all add: pochhammer_rec' mult_ac) | |
| 353 | also have "?A = (z + of_nat (Suc (2 * n + 1)) / 2) * (z + of_nat (Suc (Suc (2 * n + 1))) / 2)" | |
| 354 | (is "_ = ?B") | |
| 355 | by (simp add: field_simps n'_def) | |
| 356 | also note Suc[folded n'_def] | |
| 357 | also have "(\<Prod>k=0..2 * n + 1. z + of_nat k / 2) * ?B = (\<Prod>k=0..2 * Suc n + 1. z + of_nat k / 2)" | |
| 358 | by (simp add: atLeast0_atMost_Suc) | |
| 359 | finally show ?case | |
| 360 | by (simp add: n'_def) | |
| 361 | qed | |
| 362 | ||
| 363 | lemma pochhammer_double: | |
| 364 | fixes z :: "'a::field_char_0" | |
| 365 | shows "pochhammer (2 * z) (2 * n) = of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n" | |
| 366 | proof (induct n) | |
| 367 | case 0 | |
| 368 | then show ?case by simp | |
| 369 | next | |
| 370 | case (Suc n) | |
| 371 | have "pochhammer (2 * z) (2 * (Suc n)) = pochhammer (2 * z) (2 * n) * | |
| 372 | (2 * (z + of_nat n)) * (2 * (z + of_nat n) + 1)" | |
| 373 | by (simp add: pochhammer_rec' ac_simps) | |
| 374 | also note Suc | |
| 375 | also have "of_nat (2 ^ (2 * n)) * pochhammer z n * pochhammer (z + 1/2) n * | |
| 376 | (2 * (z + of_nat n)) * (2 * (z + of_nat n) + 1) = | |
| 377 | of_nat (2 ^ (2 * (Suc n))) * pochhammer z (Suc n) * pochhammer (z + 1/2) (Suc n)" | |
| 378 | by (simp add: field_simps pochhammer_rec') | |
| 379 | finally show ?case . | |
| 380 | qed | |
| 381 | ||
| 382 | lemma fact_double: | |
| 383 | "fact (2 * n) = (2 ^ (2 * n) * pochhammer (1 / 2) n * fact n :: 'a::field_char_0)" | |
| 384 | using pochhammer_double[of "1/2::'a" n] by (simp add: pochhammer_fact) | |
| 385 | ||
| 386 | lemma pochhammer_absorb_comp: "(r - of_nat k) * pochhammer (- r) k = r * pochhammer (-r + 1) k" | |
| 387 | (is "?lhs = ?rhs") | |
| 388 | for r :: "'a::comm_ring_1" | |
| 389 | proof - | |
| 390 | have "?lhs = - pochhammer (- r) (Suc k)" | |
| 391 | by (subst pochhammer_rec') (simp add: algebra_simps) | |
| 392 | also have "\<dots> = ?rhs" | |
| 393 | by (subst pochhammer_rec) simp | |
| 394 | finally show ?thesis . | |
| 395 | qed | |
| 396 | ||
| 397 | ||
| 398 | subsection \<open>Misc\<close> | |
| 399 | ||
| 400 | lemma fact_code [code]: | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68783diff
changeset | 401 | "fact n = (of_nat (fold_atLeastAtMost_nat ((*)) 2 n 1) :: 'a::semiring_char_0)" | 
| 65812 | 402 | proof - | 
| 403 |   have "fact n = (of_nat (\<Prod>{1..n}) :: 'a)"
 | |
| 404 | by (simp add: fact_prod) | |
| 405 |   also have "\<Prod>{1..n} = \<Prod>{2..n}"
 | |
| 406 | by (intro prod.mono_neutral_right) auto | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68783diff
changeset | 407 | also have "\<dots> = fold_atLeastAtMost_nat ((*)) 2 n 1" | 
| 65812 | 408 | by (simp add: prod_atLeastAtMost_code) | 
| 409 | finally show ?thesis . | |
| 410 | qed | |
| 411 | ||
| 412 | lemma pochhammer_code [code]: | |
| 413 | "pochhammer a n = | |
| 414 | (if n = 0 then 1 | |
| 415 | else fold_atLeastAtMost_nat (\<lambda>n acc. (a + of_nat n) * acc) 0 (n - 1) 1)" | |
| 416 | by (cases n) | |
| 417 | (simp_all add: pochhammer_prod prod_atLeastAtMost_code [symmetric] | |
| 418 | atLeastLessThanSuc_atLeastAtMost) | |
| 419 | ||
| 72569 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 420 | |
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 421 | lemma mult_less_iff1: "0 < z \<Longrightarrow> x * z < y * z \<longleftrightarrow> x < y" | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 422 | for x y z :: "'a::linordered_idom" | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 423 | by simp | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 424 | |
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 425 | lemma mult_le_cancel_iff1: "0 < z \<Longrightarrow> x * z \<le> y * z \<longleftrightarrow> x \<le> y" | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 426 | for x y z :: "'a::linordered_idom" | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 427 | by simp | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 428 | |
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 429 | lemma mult_le_cancel_iff2: "0 < z \<Longrightarrow> z * x \<le> z * y \<longleftrightarrow> x \<le> y" | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 430 | for x y z :: "'a::linordered_idom" | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 431 | by simp | 
| 
d56e4eeae967
mult_le_cancel_iff1, mult_le_cancel_iff2, mult_less_iff1 generalised from the real_ versions
 paulson <lp15@cam.ac.uk> parents: 
70113diff
changeset | 432 | |
| 65812 | 433 | end |