# HG changeset patch # User paulson # Date 1426000840 0 # Node ID 651ea265d56811ec3b7ed20be0ee282dcae52d0a # Parent 37adca7fd48ff367a116f2ddaac77a619f055c56 Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy diff -r 37adca7fd48f -r 651ea265d568 src/Doc/Tutorial/Sets/Examples.thy --- a/src/Doc/Tutorial/Sets/Examples.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/Doc/Tutorial/Sets/Examples.thy Tue Mar 10 15:20:40 2015 +0000 @@ -1,4 +1,4 @@ -theory Examples imports "~~/src/HOL/Number_Theory/Binomial" begin +theory Examples imports Complex_Main begin declare [[eta_contract = false]] diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Algebra/Exponent.thy --- a/src/HOL/Algebra/Exponent.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Algebra/Exponent.thy Tue Mar 10 15:20:40 2015 +0000 @@ -6,7 +6,7 @@ *) theory Exponent -imports Main "~~/src/HOL/Number_Theory/Primes" "~~/src/HOL/Number_Theory/Binomial" +imports Main "~~/src/HOL/Number_Theory/Primes" begin section {*Sylow's Theorem*} @@ -35,7 +35,7 @@ lemma prime_dvd_cases: fixes p::nat - shows "[| p*k dvd m*n; prime p |] + shows "[| p*k dvd m*n; prime p |] ==> (\x. k dvd x*n & m = p*x) | (\y. k dvd m*y & n = p*y)" apply (simp add: prime_iff) apply (frule dvd_mult_left) @@ -48,10 +48,10 @@ done -lemma prime_power_dvd_cases [rule_format (no_asm)]: +lemma prime_power_dvd_cases [rule_format (no_asm)]: fixes p::nat shows "prime p - ==> \m n. p^c dvd m*n --> + ==> \m n. p^c dvd m*n --> (\a b. a+b = Suc c --> p^a dvd m | p^b dvd n)" apply (induct c) apply (metis dvd_1_left nat_power_eq_Suc_0_iff one_is_add) @@ -119,7 +119,7 @@ lemma power_Suc_exponent_Not_dvd: "[|(p * p ^ exponent p s) dvd s; prime p |] ==> s=0" apply (subgoal_tac "p ^ Suc (exponent p s) dvd s") - prefer 2 apply simp + prefer 2 apply simp apply (rule ccontr) apply (drule exponent_ge, auto) done @@ -147,7 +147,7 @@ by (metis mult_dvd_mono power_exponent_dvd) (* exponent_mult_add, opposite inclusion *) -lemma exponent_mult_add2: "[| a > 0; b > 0 |] +lemma exponent_mult_add2: "[| a > 0; b > 0 |] ==> exponent p (a * b) <= (exponent p a) + (exponent p b)" apply (case_tac "prime p") apply (rule leI, clarify) @@ -155,7 +155,7 @@ apply (subgoal_tac "p ^ (Suc (exponent p a + exponent p b)) dvd a * b") apply (rule_tac [2] le_imp_power_dvd [THEN dvd_trans]) prefer 3 apply assumption - prefer 2 apply simp + prefer 2 apply simp apply (frule_tac a = "Suc (exponent p a) " and b = "Suc (exponent p b) " in prime_power_dvd_cases) apply (assumption, force, simp) apply (blast dest: power_Suc_exponent_Not_dvd) @@ -185,7 +185,7 @@ text{*Main Combinatorial Argument*} lemma gcd_mult': fixes a::nat shows "gcd b (a * b) = b" -by (simp add: mult.commute[of a b]) +by (simp add: mult.commute[of a b]) lemma le_extend_mult: "[| c > 0; a <= b |] ==> a <= b * (c::nat)" apply (rule_tac P = "%x. x <= b * c" in subst) @@ -204,7 +204,7 @@ apply (metis diff_is_0_eq dvd_diffD1 gcd_dvd2_nat gcd_mult' gr0I le_extend_mult less_diff_conv nat_dvd_not_less mult.commute not_add_less2 xt1(10)) done -lemma p_fac_forw: "[| (m::nat) > 0; k>0; k < p^a; (p^r) dvd (p^a)* m - k |] +lemma p_fac_forw: "[| (m::nat) > 0; k>0; k < p^a; (p^r) dvd (p^a)* m - k |] ==> (p^r) dvd (p^a) - k" apply (frule p_fac_forw_lemma [THEN le_imp_power_dvd, of _ k p], auto) apply (subgoal_tac "p^r dvd p^a*m") @@ -220,7 +220,7 @@ "[| (k::nat) > 0; k < p^a; p>0; (p^r) dvd (p^a) - k |] ==> r <= a" by (rule_tac m = "Suc 0" in p_fac_forw_lemma, auto) -lemma p_fac_backw: "[| m>0; k>0; (p::nat)\0; k < p^a; (p^r) dvd p^a - k |] +lemma p_fac_backw: "[| m>0; k>0; (p::nat)\0; k < p^a; (p^r) dvd p^a - k |] ==> (p^r) dvd (p^a)*m - k" apply (frule_tac k1 = k and p1 = p in r_le_a_forw [THEN le_imp_power_dvd], auto) apply (subgoal_tac "p^r dvd p^a*m") @@ -231,7 +231,7 @@ apply (drule less_imp_Suc_add, auto) done -lemma exponent_p_a_m_k_equation: "[| m>0; k>0; (p::nat)\0; k < p^a |] +lemma exponent_p_a_m_k_equation: "[| m>0; k>0; (p::nat)\0; k < p^a |] ==> exponent p (p^a * m - k) = exponent p (p^a - k)" apply (blast intro: exponent_equalityI p_fac_forw p_fac_backw) done @@ -241,16 +241,16 @@ (*The bound K is needed; otherwise it's too weak to be used.*) lemma p_not_div_choose_lemma [rule_format]: - "[| \i. Suc i < K --> exponent p (Suc i) = exponent p (Suc(j+i))|] + "[| \i. Suc i < K --> exponent p (Suc i) = exponent p (Suc(j+i))|] ==> k exponent p ((j+k) choose k) = 0" apply (cases "prime p") - prefer 2 apply simp + prefer 2 apply simp apply (induct k) apply (simp (no_asm)) (*induction step*) apply (subgoal_tac "(Suc (j+k) choose Suc k) > 0") prefer 2 apply (simp, clarify) -apply (subgoal_tac "exponent p ((Suc (j+k) choose Suc k) * Suc k) = +apply (subgoal_tac "exponent p ((Suc (j+k) choose Suc k) * Suc k) = exponent p (Suc k)") txt{*First, use the assumed equation. We simplify the LHS to @{term "exponent p (Suc (j + k) choose Suc k) + exponent p (Suc k)"} @@ -276,7 +276,7 @@ lemma const_p_fac_right: "m>0 ==> exponent p ((p^a * m - Suc 0) choose (p^a - Suc 0)) = 0" apply (case_tac "prime p") - prefer 2 apply simp + prefer 2 apply simp apply (frule_tac a = a in zero_less_prime_power) apply (rule_tac K = "p^a" in p_not_div_choose) apply simp @@ -294,14 +294,14 @@ lemma const_p_fac: "m>0 ==> exponent p (((p^a) * m) choose p^a) = exponent p m" apply (case_tac "prime p") - prefer 2 apply simp + prefer 2 apply simp apply (subgoal_tac "0 < p^a * m & p^a <= p^a * m") prefer 2 apply (force simp add: prime_iff) txt{*A similar trick to the one used in @{text p_not_div_choose_lemma}: insert an equation; use @{text exponent_mult_add} on the LHS; on the RHS, first transform the binomial coefficient, then use @{text exponent_mult_add}.*} -apply (subgoal_tac "exponent p ((( (p^a) * m) choose p^a) * p^a) = +apply (subgoal_tac "exponent p ((( (p^a) * m) choose p^a) * p^a) = a + exponent p m") apply (simp add: exponent_mult_add) txt{*one subgoal left!*} diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Fact.thy --- a/src/HOL/Fact.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Fact.thy Tue Mar 10 15:20:40 2015 +0000 @@ -15,7 +15,7 @@ fixes fact :: "'a \ 'a" instantiation nat :: fact -begin +begin fun fact_nat :: "nat \ nat" @@ -31,11 +31,11 @@ instantiation int :: fact -begin +begin definition fact_int :: "int \ int" -where +where "fact_int x = (if x >= 0 then int (fact (nat x)) else 0)" instance proof qed @@ -55,7 +55,7 @@ "x >= (0::int) \ fact x >= 0" by (auto simp add: fact_int_def) -declare transfer_morphism_nat_int[transfer add return: +declare transfer_morphism_nat_int[transfer add return: transfer_nat_int_factorial transfer_nat_int_factorial_closure] lemma transfer_int_nat_factorial: @@ -66,7 +66,7 @@ "is_nat x \ fact x >= 0" by (auto simp add: fact_int_def) -declare transfer_morphism_int_nat[transfer add return: +declare transfer_morphism_int_nat[transfer add return: transfer_int_nat_factorial transfer_int_nat_factorial_closure] @@ -87,10 +87,10 @@ lemma fact_plus_one_nat: "fact ((n::nat) + 1) = (n + 1) * fact n" by simp -lemma fact_plus_one_int: +lemma fact_plus_one_int: assumes "n >= 0" shows "fact ((n::int) + 1) = (n + 1) * fact n" - using assms unfolding fact_int_def + using assms unfolding fact_int_def by (simp add: nat_add_distrib algebra_simps int_mult) lemma fact_reduce_nat: "(n::nat) > 0 \ fact n = n * fact (n - 1)" @@ -153,7 +153,7 @@ apply auto done -lemma interval_plus_one_nat: "(i::nat) <= j + 1 \ +lemma interval_plus_one_nat: "(i::nat) <= j + 1 \ {i..j+1} = {i..j} Un {j+1}" by auto @@ -199,7 +199,7 @@ case (Suc d') have "fact (n + Suc d') div fact n = Suc (n + d') * fact (n + d') div fact n" by simp - also from Suc.hyps have "... = Suc (n + d') * \{n + 1..n + d'}" + also from Suc.hyps have "... = Suc (n + d') * \{n + 1..n + d'}" unfolding div_mult1_eq[of _ "fact (n + d')"] by (simp add: fact_mod) also have "... = \{n + 1..n + Suc d'}" by (simp add: atLeastAtMostSuc_conv setprod.insert) @@ -224,7 +224,7 @@ apply arith done -lemma fact_mono_int_aux [rule_format]: "k >= (0::int) \ +lemma fact_mono_int_aux [rule_format]: "k >= (0::int) \ fact (m + k) >= fact m" apply (case_tac "m < 0") apply auto @@ -266,7 +266,7 @@ apply auto done -lemma fact_num_eq_if_nat: "fact (m::nat) = +lemma fact_num_eq_if_nat: "fact (m::nat) = (if m=0 then 1 else m * fact (m - 1))" by (cases m) auto @@ -275,7 +275,7 @@ by (cases "m + n") auto lemma fact_add_num_eq_if2_nat: - "fact ((m::nat) + n) = + "fact ((m::nat) + n) = (if m = 0 then fact n else (m + n) * fact ((m - 1) + n))" by (cases m) auto @@ -339,7 +339,7 @@ lemma binomial_Suc_Suc [simp]: "(Suc n choose Suc k) = (n choose k) + (n choose Suc k)" by simp -lemma choose_reduce_nat: +lemma choose_reduce_nat: "0 < (n::nat) \ 0 < k \ (n choose k) = ((n - 1) choose (k - 1)) + ((n - 1) choose k)" by (metis Suc_diff_1 binomial.simps(2) neq0_conv) @@ -404,7 +404,7 @@ {s. s \ M \ card s = Suc k} \ {s. \t. t \ M \ card t = k \ s = insert x t}" apply safe apply (auto intro: finite_subset [THEN card_insert_disjoint]) - by (metis (full_types) Diff_insert_absorb Set.set_insert Zero_neq_Suc card_Diff_singleton_if + by (metis (full_types) Diff_insert_absorb Set.set_insert Zero_neq_Suc card_Diff_singleton_if card_eq_0_iff diff_Suc_1 in_mono subset_insert_iff) lemma finite_bex_subset [simp]: @@ -455,7 +455,7 @@ subsection {* The binomial theorem (courtesy of Tobias Nipkow): *} text{* Avigad's version, generalized to any commutative ring *} -theorem binomial_ring: "(a+b::'a::{comm_ring_1,power})^n = +theorem binomial_ring: "(a+b::'a::{comm_ring_1,power})^n = (\k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))" (is "?P n") proof (induct n) case 0 then show "?P 0" by simp @@ -465,7 +465,7 @@ by auto have decomp2: "{0..n} = {0} Un {1..n}" by auto - have "(a+b)^(n+1) = + have "(a+b)^(n+1) = (a+b) * (\k=0..n. of_nat (n choose k) * a^k * b^(n-k))" using Suc.hyps by simp also have "\ = a*(\k=0..n. of_nat (n choose k) * a^k * b^(n-k)) + @@ -476,14 +476,14 @@ by (auto simp add: setsum_right_distrib ac_simps) also have "\ = (\k=0..n. of_nat (n choose k) * a^k * b^(n+1-k)) + (\k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n+1-k))" - by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le field_simps + by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le field_simps del:setsum_cl_ivl_Suc) also have "\ = a^(n+1) + b^(n+1) + (\k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n+1-k)) + (\k=1..n. of_nat (n choose k) * a^k * b^(n+1-k))" by (simp add: decomp2) also have - "\ = a^(n+1) + b^(n+1) + + "\ = a^(n+1) + b^(n+1) + (\k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))" by (auto simp add: field_simps setsum.distrib [symmetric] choose_reduce_nat) also have "\ = (\k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))" @@ -518,7 +518,7 @@ by simp from n h th0 have "fact k * fact (n - k) * (n choose k) = - k * (fact h * fact (m - h) * (m choose h)) + + k * (fact h * fact (m - h) * (m choose h)) + (m - h) * (fact k * fact (m - k) * (m choose k))" by (simp add: field_simps) also have "\ = (k + (m - h)) * fact m" @@ -537,4 +537,644 @@ using binomial_fact_lemma[OF kn] by (simp add: field_simps of_nat_mult [symmetric]) +lemma choose_row_sum: "(\k=0..n. n choose k) = 2^n" + using binomial [of 1 "1" n] + by (simp add: numeral_2_eq_2) + +lemma sum_choose_lower: "(\k=0..n. (r+k) choose k) = Suc (r+n) choose n" + by (induct n) auto + +lemma sum_choose_upper: "(\k=0..n. k choose m) = Suc n choose Suc m" + by (induct n) auto + +lemma natsum_reverse_index: + fixes m::nat + shows "(\k. m \ k \ k \ n \ g k = f (m + n - k)) \ (\k=m..n. f k) = (\k=m..n. g k)" + by (rule setsum.reindex_bij_witness[where i="\k. m+n-k" and j="\k. m+n-k"]) auto + +text{*NW diagonal sum property*} +lemma sum_choose_diagonal: + assumes "m\n" shows "(\k=0..m. (n-k) choose (m-k)) = Suc n choose m" +proof - + have "(\k=0..m. (n-k) choose (m-k)) = (\k=0..m. (n-m+k) choose k)" + by (rule natsum_reverse_index) (simp add: assms) + also have "... = Suc (n-m+m) choose m" + by (rule sum_choose_lower) + also have "... = Suc n choose m" using assms + by simp + finally show ?thesis . +qed + +subsection{* Pochhammer's symbol : generalized rising factorial *} + +text {* See @{url "http://en.wikipedia.org/wiki/Pochhammer_symbol"} *} + +definition "pochhammer (a::'a::comm_semiring_1) n = + (if n = 0 then 1 else setprod (\n. a + of_nat n) {0 .. n - 1})" + +lemma pochhammer_0 [simp]: "pochhammer a 0 = 1" + by (simp add: pochhammer_def) + +lemma pochhammer_1 [simp]: "pochhammer a 1 = a" + by (simp add: pochhammer_def) + +lemma pochhammer_Suc0 [simp]: "pochhammer a (Suc 0) = a" + by (simp add: pochhammer_def) + +lemma pochhammer_Suc_setprod: "pochhammer a (Suc n) = setprod (\n. a + of_nat n) {0 .. n}" + by (simp add: pochhammer_def) + +lemma setprod_nat_ivl_Suc: "setprod f {0 .. Suc n} = setprod f {0..n} * f (Suc n)" +proof - + have "{0..Suc n} = {0..n} \ {Suc n}" by auto + then show ?thesis by (simp add: field_simps) +qed + +lemma setprod_nat_ivl_1_Suc: "setprod f {0 .. Suc n} = f 0 * setprod f {1.. Suc n}" +proof - + have "{0..Suc n} = {0} \ {1 .. Suc n}" by auto + then show ?thesis by simp +qed + + +lemma pochhammer_Suc: "pochhammer a (Suc n) = pochhammer a n * (a + of_nat n)" +proof (cases n) + case 0 + then show ?thesis by simp +next + case (Suc n) + show ?thesis unfolding Suc pochhammer_Suc_setprod setprod_nat_ivl_Suc .. +qed + +lemma pochhammer_rec: "pochhammer a (Suc n) = a * pochhammer (a + 1) n" +proof (cases "n = 0") + case True + then show ?thesis by (simp add: pochhammer_Suc_setprod) +next + case False + have *: "finite {1 .. n}" "0 \ {1 .. n}" by auto + have eq: "insert 0 {1 .. n} = {0..n}" by auto + have **: "(\n\{1\nat..n}. a + of_nat n) = (\n\{0\nat..n - 1}. a + 1 + of_nat n)" + apply (rule setprod.reindex_cong [where l = Suc]) + using False + apply (auto simp add: fun_eq_iff field_simps) + done + show ?thesis + apply (simp add: pochhammer_def) + unfolding setprod.insert [OF *, unfolded eq] + using ** apply (simp add: field_simps) + done +qed + +lemma pochhammer_fact: "of_nat (fact n) = pochhammer 1 n" + unfolding fact_altdef_nat + apply (cases n) + apply (simp_all add: of_nat_setprod pochhammer_Suc_setprod) + apply (rule setprod.reindex_cong [where l = Suc]) + apply (auto simp add: fun_eq_iff) + done + +lemma pochhammer_of_nat_eq_0_lemma: + assumes "k > n" + shows "pochhammer (- (of_nat n :: 'a:: idom)) k = 0" +proof (cases "n = 0") + case True + then show ?thesis + using assms by (cases k) (simp_all add: pochhammer_rec) +next + case False + from assms obtain h where "k = Suc h" by (cases k) auto + then show ?thesis + by (simp add: pochhammer_Suc_setprod) + (metis Suc_leI Suc_le_mono assms atLeastAtMost_iff less_eq_nat.simps(1)) +qed + +lemma pochhammer_of_nat_eq_0_lemma': + assumes kn: "k \ n" + shows "pochhammer (- (of_nat n :: 'a:: {idom,ring_char_0})) k \ 0" +proof (cases k) + case 0 + then show ?thesis by simp +next + case (Suc h) + then show ?thesis + apply (simp add: pochhammer_Suc_setprod) + using Suc kn apply (auto simp add: algebra_simps) + done +qed + +lemma pochhammer_of_nat_eq_0_iff: + shows "pochhammer (- (of_nat n :: 'a:: {idom,ring_char_0})) k = 0 \ k > n" + (is "?l = ?r") + using pochhammer_of_nat_eq_0_lemma[of n k, where ?'a='a] + pochhammer_of_nat_eq_0_lemma'[of k n, where ?'a = 'a] + by (auto simp add: not_le[symmetric]) + +lemma pochhammer_eq_0_iff: "pochhammer a n = (0::'a::field_char_0) \ (\k < n. a = - of_nat k)" + apply (auto simp add: pochhammer_of_nat_eq_0_iff) + apply (cases n) + apply (auto simp add: pochhammer_def algebra_simps group_add_class.eq_neg_iff_add_eq_0) + apply (metis leD not_less_eq) + done + +lemma pochhammer_eq_0_mono: + "pochhammer a n = (0::'a::field_char_0) \ m \ n \ pochhammer a m = 0" + unfolding pochhammer_eq_0_iff by auto + +lemma pochhammer_neq_0_mono: + "pochhammer a m \ (0::'a::field_char_0) \ m \ n \ pochhammer a n \ 0" + unfolding pochhammer_eq_0_iff by auto + +lemma pochhammer_minus: + assumes kn: "k \ n" + shows "pochhammer (- b) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (b - of_nat k + 1) k" +proof (cases k) + case 0 + then show ?thesis by simp +next + case (Suc h) + have eq: "((- 1) ^ Suc h :: 'a) = (\i=0..h. - 1)" + using setprod_constant[where A="{0 .. h}" and y="- 1 :: 'a"] + by auto + show ?thesis + unfolding Suc pochhammer_Suc_setprod eq setprod.distrib[symmetric] + by (rule setprod.reindex_bij_witness[where i="op - h" and j="op - h"]) + (auto simp: of_nat_diff) +qed + +lemma pochhammer_minus': + assumes kn: "k \ n" + shows "pochhammer (b - of_nat k + 1) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (- b) k" + unfolding pochhammer_minus[OF kn, where b=b] + unfolding mult.assoc[symmetric] + unfolding power_add[symmetric] + by simp + +lemma pochhammer_same: "pochhammer (- of_nat n) n = + ((- 1) ^ n :: 'a::comm_ring_1) * of_nat (fact n)" + unfolding pochhammer_minus[OF le_refl[of n]] + by (simp add: of_nat_diff pochhammer_fact) + + +subsection{* Generalized binomial coefficients *} + +definition gbinomial :: "'a::field_char_0 \ nat \ 'a" (infixl "gchoose" 65) + where "a gchoose n = + (if n = 0 then 1 else (setprod (\i. a - of_nat i) {0 .. n - 1}) / of_nat (fact n))" + +lemma gbinomial_0 [simp]: "a gchoose 0 = 1" "0 gchoose (Suc n) = 0" + apply (simp_all add: gbinomial_def) + apply (subgoal_tac "(\i\nat\{0\nat..n}. - of_nat i) = (0::'b)") + apply (simp del:setprod_zero_iff) + apply simp + done + +lemma gbinomial_pochhammer: "a gchoose n = (- 1) ^ n * pochhammer (- a) n / of_nat (fact n)" +proof (cases "n = 0") + case True + then show ?thesis by simp +next + case False + from this setprod_constant[of "{0 .. n - 1}" "- (1:: 'a)"] + have eq: "(- (1\'a)) ^ n = setprod (\i. - 1) {0 .. n - 1}" + by auto + from False show ?thesis + by (simp add: pochhammer_def gbinomial_def field_simps + eq setprod.distrib[symmetric]) +qed + +lemma binomial_gbinomial: "of_nat (n choose k) = of_nat n gchoose k" +proof - + { assume kn: "k > n" + then have ?thesis + by (subst binomial_eq_0[OF kn]) + (simp add: gbinomial_pochhammer field_simps pochhammer_of_nat_eq_0_iff) } + moreover + { assume "k=0" then have ?thesis by simp } + moreover + { assume kn: "k \ n" and k0: "k\ 0" + from k0 obtain h where h: "k = Suc h" by (cases k) auto + from h + have eq:"(- 1 :: 'a) ^ k = setprod (\i. - 1) {0..h}" + by (subst setprod_constant) auto + have eq': "(\i\{0..h}. of_nat n + - (of_nat i :: 'a)) = (\i\{n - h..n}. of_nat i)" + using h kn + by (intro setprod.reindex_bij_witness[where i="op - n" and j="op - n"]) + (auto simp: of_nat_diff) + have th0: "finite {1..n - Suc h}" "finite {n - h .. n}" + "{1..n - Suc h} \ {n - h .. n} = {}" and + eq3: "{1..n - Suc h} \ {n - h .. n} = {1..n}" + using h kn by auto + from eq[symmetric] + have ?thesis using kn + apply (simp add: binomial_fact[OF kn, where ?'a = 'a] + gbinomial_pochhammer field_simps pochhammer_Suc_setprod) + apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h + of_nat_setprod setprod.distrib[symmetric] eq' del: One_nat_def power_Suc) + unfolding setprod.union_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \ 'a"] eq[unfolded h] + unfolding mult.assoc[symmetric] + unfolding setprod.distrib[symmetric] + apply simp + apply (intro setprod.reindex_bij_witness[where i="op - n" and j="op - n"]) + apply (auto simp: of_nat_diff) + done + } + moreover + have "k > n \ k = 0 \ (k \ n \ k \ 0)" by arith + ultimately show ?thesis by blast +qed + +lemma gbinomial_1[simp]: "a gchoose 1 = a" + by (simp add: gbinomial_def) + +lemma gbinomial_Suc0[simp]: "a gchoose (Suc 0) = a" + by (simp add: gbinomial_def) + +lemma gbinomial_mult_1: + "a * (a gchoose n) = + of_nat n * (a gchoose n) + of_nat (Suc n) * (a gchoose (Suc n))" (is "?l = ?r") +proof - + have "?r = ((- 1) ^n * pochhammer (- a) n / of_nat (fact n)) * (of_nat n - (- a + of_nat n))" + unfolding gbinomial_pochhammer + pochhammer_Suc fact_Suc of_nat_mult right_diff_distrib power_Suc + by (simp add: field_simps del: of_nat_Suc) + also have "\ = ?l" unfolding gbinomial_pochhammer + by (simp add: field_simps) + finally show ?thesis .. +qed + +lemma gbinomial_mult_1': + "(a gchoose n) * a = of_nat n * (a gchoose n) + of_nat (Suc n) * (a gchoose (Suc n))" + by (simp add: mult.commute gbinomial_mult_1) + +lemma gbinomial_Suc: + "a gchoose (Suc k) = (setprod (\i. a - of_nat i) {0 .. k}) / of_nat (fact (Suc k))" + by (simp add: gbinomial_def) + +lemma gbinomial_mult_fact: + "(of_nat (fact (Suc k)) :: 'a) * ((a::'a::field_char_0) gchoose (Suc k)) = + (setprod (\i. a - of_nat i) {0 .. k})" + by (simp_all add: gbinomial_Suc field_simps del: fact_Suc) + +lemma gbinomial_mult_fact': + "((a::'a::field_char_0) gchoose (Suc k)) * (of_nat (fact (Suc k)) :: 'a) = + (setprod (\i. a - of_nat i) {0 .. k})" + using gbinomial_mult_fact[of k a] + by (subst mult.commute) + + +lemma gbinomial_Suc_Suc: + "((a::'a::field_char_0) + 1) gchoose (Suc k) = a gchoose k + (a gchoose (Suc k))" +proof (cases k) + case 0 + then show ?thesis by simp +next + case (Suc h) + have eq0: "(\i\{1..k}. (a + 1) - of_nat i) = (\i\{0..h}. a - of_nat i)" + apply (rule setprod.reindex_cong [where l = Suc]) + using Suc + apply auto + done + have "of_nat (fact (Suc k)) * (a gchoose k + (a gchoose (Suc k))) = + ((a gchoose Suc h) * of_nat (fact (Suc h)) * of_nat (Suc k)) + (\i\{0\nat..Suc h}. a - of_nat i)" + apply (simp add: Suc field_simps del: fact_Suc) + unfolding gbinomial_mult_fact' + apply (subst fact_Suc) + unfolding of_nat_mult + apply (subst mult.commute) + unfolding mult.assoc + unfolding gbinomial_mult_fact + apply (simp add: field_simps) + done + also have "\ = (\i\{0..h}. a - of_nat i) * (a + 1)" + unfolding gbinomial_mult_fact' setprod_nat_ivl_Suc + by (simp add: field_simps Suc) + also have "\ = (\i\{0..k}. (a + 1) - of_nat i)" + using eq0 + by (simp add: Suc setprod_nat_ivl_1_Suc) + also have "\ = of_nat (fact (Suc k)) * ((a + 1) gchoose (Suc k))" + unfolding gbinomial_mult_fact .. + finally show ?thesis by (simp del: fact_Suc) +qed + +lemma gbinomial_reduce_nat: + "0 < k \ (a::'a::field_char_0) gchoose k = (a - 1) gchoose (k - 1) + ((a - 1) gchoose k)" +by (metis Suc_pred' diff_add_cancel gbinomial_Suc_Suc) + + +lemma binomial_symmetric: + assumes kn: "k \ n" + shows "n choose k = n choose (n - k)" +proof- + from kn have kn': "n - k \ n" by arith + from binomial_fact_lemma[OF kn] binomial_fact_lemma[OF kn'] + have "fact k * fact (n - k) * (n choose k) = + fact (n - k) * fact (n - (n - k)) * (n choose (n - k))" by simp + then show ?thesis using kn by simp +qed + +text{*Contributed by Manuel Eberl, generalised by LCP. + Alternative definition of the binomial coefficient as @{term "\iii = (\i x" + shows "(x / of_nat k :: 'a) ^ k \ x gchoose k" +proof - + have x: "0 \ x" + using assms of_nat_0_le_iff order_trans by blast + have "(x / of_nat k :: 'a) ^ k = (\i \ x gchoose k" + unfolding gbinomial_altdef_of_nat + proof (safe intro!: setprod_mono) + fix i :: nat + assume ik: "i < k" + from assms have "x * of_nat i \ of_nat (i * k)" + by (metis mult.commute mult_le_cancel_right of_nat_less_0_iff of_nat_mult) + then have "x * of_nat k - x * of_nat i \ x * of_nat k - of_nat (i * k)" by arith + then have "x * of_nat (k - i) \ (x - of_nat i) * of_nat k" + using ik + by (simp add: algebra_simps zero_less_mult_iff of_nat_diff of_nat_mult) + then have "x * of_nat (k - i) \ (x - of_nat i) * (of_nat k :: 'a)" + unfolding of_nat_mult[symmetric] of_nat_le_iff . + with assms show "x / of_nat k \ (x - of_nat i) / (of_nat (k - i) :: 'a)" + using `i < k` by (simp add: field_simps) + qed (simp add: x zero_le_divide_iff) + finally show ?thesis . +qed + +text{*Versions of the theorems above for the natural-number version of "choose"*} +lemma binomial_altdef_of_nat: + fixes n k :: nat + and x :: "'a :: {field_char_0,field_inverse_zero}" --{*the point is to constrain @{typ 'a}*} + assumes "k \ n" + shows "of_nat (n choose k) = (\i n" + shows "(of_nat n / of_nat k :: 'a) ^ k \ of_nat (n choose k)" +by (simp add: assms gbinomial_ge_n_over_k_pow_k binomial_gbinomial of_nat_diff) + +lemma binomial_le_pow: + assumes "r \ n" + shows "n choose r \ n ^ r" +proof - + have "n choose r \ fact n div fact (n - r)" + using `r \ n` by (subst binomial_fact_lemma[symmetric]) auto + with fact_div_fact_le_pow [OF assms] show ?thesis by auto +qed + +lemma binomial_altdef_nat: "(k::nat) \ n \ + n choose k = fact n div (fact k * fact (n - k))" + by (subst binomial_fact_lemma [symmetric]) auto + +lemma choose_dvd_nat: "(k::nat) \ n \ fact k * fact (n - k) dvd fact n" +by (metis binomial_fact_lemma dvd_def) + +lemma choose_dvd_int: + assumes "(0::int) <= k" and "k <= n" + shows "fact k * fact (n - k) dvd fact n" + apply (subst tsub_eq [symmetric], rule assms) + apply (rule choose_dvd_nat [transferred]) + using assms apply auto + done + +lemma fact_fact_dvd_fact: fixes k::nat shows "fact k * fact n dvd fact (n + k)" +by (metis add.commute add_diff_cancel_left' choose_dvd_nat le_add2) + +lemma choose_mult_lemma: + "((m+r+k) choose (m+k)) * ((m+k) choose k) = ((m+r+k) choose k) * ((m+r) choose m)" +proof - + have "((m+r+k) choose (m+k)) * ((m+k) choose k) = + fact (m+r + k) div (fact (m + k) * fact (m+r - m)) * (fact (m + k) div (fact k * fact m))" + by (simp add: assms binomial_altdef_nat) + also have "... = fact (m+r+k) div (fact r * (fact k * fact m))" + apply (subst div_mult_div_if_dvd) + apply (auto simp: fact_fact_dvd_fact) + apply (metis add.assoc add.commute fact_fact_dvd_fact) + done + also have "... = (fact (m+r+k) * fact (m+r)) div (fact r * (fact k * fact m) * fact (m+r))" + apply (subst div_mult_div_if_dvd [symmetric]) + apply (auto simp: fact_fact_dvd_fact) + apply (metis dvd_trans dvd.dual_order.refl fact_fact_dvd_fact mult_dvd_mono mult.left_commute) + done + also have "... = (fact (m+r+k) div (fact k * fact (m+r)) * (fact (m+r) div (fact r * fact m)))" + apply (subst div_mult_div_if_dvd) + apply (auto simp: fact_fact_dvd_fact) + apply(metis mult.left_commute) + done + finally show ?thesis + by (simp add: binomial_altdef_nat mult.commute) +qed + +text{*The "Subset of a Subset" identity*} +lemma choose_mult: + assumes "k\m" "m\n" + shows "(n choose m) * (m choose k) = (n choose k) * ((n-k) choose (m-k))" +using assms choose_mult_lemma [of "m-k" "n-m" k] +by simp + + +subsection {* Binomial coefficients *} + +lemma choose_one: "(n::nat) choose 1 = n" + by simp + +(*FIXME: messy and apparently unused*) +lemma binomial_induct [rule_format]: "(ALL (n::nat). P n n) \ + (ALL n. P (Suc n) 0) \ (ALL n. (ALL k < n. P n k \ P n (Suc k) \ + P (Suc n) (Suc k))) \ (ALL k <= n. P n k)" + apply (induct n) + apply auto + apply (case_tac "k = 0") + apply auto + apply (case_tac "k = Suc n") + apply auto + apply (metis Suc_le_eq fact_nat.cases le_Suc_eq le_eq_less_or_eq) + done + +lemma card_UNION: + assumes "finite A" and "\k \ A. finite k" + shows "card (\A) = nat (\I | I \ A \ I \ {}. (- 1) ^ (card I + 1) * int (card (\I)))" + (is "?lhs = ?rhs") +proof - + have "?rhs = nat (\I | I \ A \ I \ {}. (- 1) ^ (card I + 1) * (\_\\I. 1))" by simp + also have "\ = nat (\I | I \ A \ I \ {}. (\_\\I. (- 1) ^ (card I + 1)))" (is "_ = nat ?rhs") + by(subst setsum_right_distrib) simp + also have "?rhs = (\(I, _)\Sigma {I. I \ A \ I \ {}} Inter. (- 1) ^ (card I + 1))" + using assms by(subst setsum.Sigma)(auto) + also have "\ = (\(x, I)\(SIGMA x:UNIV. {I. I \ A \ I \ {} \ x \ \I}). (- 1) ^ (card I + 1))" + by (rule setsum.reindex_cong [where l = "\(x, y). (y, x)"]) (auto intro: inj_onI simp add: split_beta) + also have "\ = (\(x, I)\(SIGMA x:\A. {I. I \ A \ I \ {} \ x \ \I}). (- 1) ^ (card I + 1))" + using assms by(auto intro!: setsum.mono_neutral_cong_right finite_SigmaI2 intro: finite_subset[where B="\A"]) + also have "\ = (\x\\A. (\I|I \ A \ I \ {} \ x \ \I. (- 1) ^ (card I + 1)))" + using assms by(subst setsum.Sigma) auto + also have "\ = (\_\\A. 1)" (is "setsum ?lhs _ = _") + proof(rule setsum.cong[OF refl]) + fix x + assume x: "x \ \A" + def K \ "{X \ A. x \ X}" + with `finite A` have K: "finite K" by auto + let ?I = "\i. {I. I \ A \ card I = i \ x \ \I}" + have "inj_on snd (SIGMA i:{1..card A}. ?I i)" + using assms by(auto intro!: inj_onI) + moreover have [symmetric]: "snd ` (SIGMA i:{1..card A}. ?I i) = {I. I \ A \ I \ {} \ x \ \I}" + using assms by(auto intro!: rev_image_eqI[where x="(card a, a)" for a] + simp add: card_gt_0_iff[folded Suc_le_eq] + dest: finite_subset intro: card_mono) + ultimately have "?lhs x = (\(i, I)\(SIGMA i:{1..card A}. ?I i). (- 1) ^ (i + 1))" + by (rule setsum.reindex_cong [where l = snd]) fastforce + also have "\ = (\i=1..card A. (\I|I \ A \ card I = i \ x \ \I. (- 1) ^ (i + 1)))" + using assms by(subst setsum.Sigma) auto + also have "\ = (\i=1..card A. (- 1) ^ (i + 1) * (\I|I \ A \ card I = i \ x \ \I. 1))" + by(subst setsum_right_distrib) simp + also have "\ = (\i=1..card K. (- 1) ^ (i + 1) * (\I|I \ K \ card I = i. 1))" (is "_ = ?rhs") + proof(rule setsum.mono_neutral_cong_right[rule_format]) + show "{1..card K} \ {1..card A}" using `finite A` + by(auto simp add: K_def intro: card_mono) + next + fix i + assume "i \ {1..card A} - {1..card K}" + hence i: "i \ card A" "card K < i" by auto + have "{I. I \ A \ card I = i \ x \ \I} = {I. I \ K \ card I = i}" + by(auto simp add: K_def) + also have "\ = {}" using `finite A` i + by(auto simp add: K_def dest: card_mono[rotated 1]) + finally show "(- 1) ^ (i + 1) * (\I | I \ A \ card I = i \ x \ \I. 1 :: int) = 0" + by(simp only:) simp + next + fix i + have "(\I | I \ A \ card I = i \ x \ \I. 1) = (\I | I \ K \ card I = i. 1 :: int)" + (is "?lhs = ?rhs") + by(rule setsum.cong)(auto simp add: K_def) + thus "(- 1) ^ (i + 1) * ?lhs = (- 1) ^ (i + 1) * ?rhs" by simp + qed simp + also have "{I. I \ K \ card I = 0} = {{}}" using assms + by(auto simp add: card_eq_0_iff K_def dest: finite_subset) + hence "?rhs = (\i = 0..card K. (- 1) ^ (i + 1) * (\I | I \ K \ card I = i. 1 :: int)) + 1" + by(subst (2) setsum_head_Suc)(simp_all ) + also have "\ = (\i = 0..card K. (- 1) * ((- 1) ^ i * int (card K choose i))) + 1" + using K by(subst n_subsets[symmetric]) simp_all + also have "\ = - (\i = 0..card K. (- 1) ^ i * int (card K choose i)) + 1" + by(subst setsum_right_distrib[symmetric]) simp + also have "\ = - ((-1 + 1) ^ card K) + 1" + by(subst binomial_ring)(simp add: ac_simps) + also have "\ = 1" using x K by(auto simp add: K_def card_gt_0_iff) + finally show "?lhs x = 1" . + qed + also have "nat \ = card (\A)" by simp + finally show ?thesis .. +qed + +text{* The number of nat lists of length @{text m} summing to @{text N} is +@{term "(N + m - 1) choose N"}: *} + +lemma card_length_listsum_rec: + assumes "m\1" + shows "card {l::nat list. length l = m \ listsum l = N} = + (card {l. length l = (m - 1) \ listsum l = N} + + card {l. length l = m \ listsum l + 1 = N})" + (is "card ?C = (card ?A + card ?B)") +proof - + let ?A'="{l. length l = m \ listsum l = N \ hd l = 0}" + let ?B'="{l. length l = m \ listsum l = N \ hd l \ 0}" + let ?f ="\ l. 0#l" + let ?g ="\ l. (hd l + 1) # tl l" + have 1: "\xs x. xs \ [] \ x = hd xs \ x # tl xs = xs" by simp + have 2: "\xs. (xs::nat list) \ [] \ listsum(tl xs) = listsum xs - hd xs" + by(auto simp add: neq_Nil_conv) + have f: "bij_betw ?f ?A ?A'" + apply(rule bij_betw_byWitness[where f' = tl]) + using assms + by (auto simp: 2 length_0_conv[symmetric] 1 simp del: length_0_conv) + have 3: "\xs:: nat list. xs \ [] \ hd xs + (listsum xs - hd xs) = listsum xs" + by (metis 1 listsum_simps(2) 2) + have g: "bij_betw ?g ?B ?B'" + apply(rule bij_betw_byWitness[where f' = "\ l. (hd l - 1) # tl l"]) + using assms + by (auto simp: 2 length_0_conv[symmetric] intro!: 3 + simp del: length_greater_0_conv length_0_conv) + { fix M N :: nat have "finite {xs. size xs = M \ set xs \ {0.. ?B'" by auto + have disj: "?A' \ ?B' = {}" by auto + have "card ?C = card(?A' \ ?B')" using uni by simp + also have "\ = card ?A + card ?B" + using card_Un_disjoint[OF _ _ disj] bij_betw_finite[OF f] bij_betw_finite[OF g] + bij_betw_same_card[OF f] bij_betw_same_card[OF g] fin_A fin_B + by presburger + finally show ?thesis . +qed + +lemma card_length_listsum: --"By Holden Lee, tidied by Tobias Nipkow" + "card {l::nat list. size l = m \ listsum l = N} = (N + m - 1) choose N" +proof (cases m) + case 0 then show ?thesis + by (cases N) (auto simp: cong: conj_cong) +next + case (Suc m') + have m: "m\1" by (simp add: Suc) + then show ?thesis + proof (induct "N + m - 1" arbitrary: N m) + case 0 -- "In the base case, the only solution is [0]." + have [simp]: "{l::nat list. length l = Suc 0 \ (\n\set l. n = 0)} = {[0]}" + by (auto simp: length_Suc_conv) + have "m=1 \ N=0" using 0 by linarith + then show ?case by simp + next + case (Suc k) + + have c1: "card {l::nat list. size l = (m - 1) \ listsum l = N} = + (N + (m - 1) - 1) choose N" + proof cases + assume "m = 1" + with Suc.hyps have "N\1" by auto + with `m = 1` show ?thesis by (simp add: binomial_eq_0) + next + assume "m \ 1" thus ?thesis using Suc by fastforce + qed + + from Suc have c2: "card {l::nat list. size l = m \ listsum l + 1 = N} = + (if N>0 then ((N - 1) + m - 1) choose (N - 1) else 0)" + proof - + have aux: "\m n. n > 0 \ Suc m = n \ m = n - 1" by arith + from Suc have "N>0 \ + card {l::nat list. size l = m \ listsum l + 1 = N} = + ((N - 1) + m - 1) choose (N - 1)" by (simp add: aux) + thus ?thesis by auto + qed + + from Suc.prems have "(card {l::nat list. size l = (m - 1) \ listsum l = N} + + card {l::nat list. size l = m \ listsum l + 1 = N}) = (N + m - 1) choose N" + by (auto simp: c1 c2 choose_reduce_nat[of "N + m - 1" N] simp del: One_nat_def) + thus ?case using card_length_listsum_rec[OF Suc.prems] by auto + qed +qed + end diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Fields.thy --- a/src/HOL/Fields.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Fields.thy Tue Mar 10 15:20:40 2015 +0000 @@ -91,7 +91,7 @@ apply auto done -lemma inverse_unique: +lemma inverse_unique: assumes ab: "a * b = 1" shows "inverse a = b" proof - @@ -121,7 +121,7 @@ lemma inverse_1 [simp]: "inverse 1 = 1" by (rule inverse_unique) simp -lemma nonzero_inverse_mult_distrib: +lemma nonzero_inverse_mult_distrib: assumes "a \ 0" and "b \ 0" shows "inverse (a * b) = inverse b * inverse a" proof - @@ -199,7 +199,7 @@ proof - assume [simp]: "c \ 0" have "b / c = a \ (b / c) * c = a * c" by simp - also have "... \ b = a * c" by (simp add: divide_inverse mult.assoc) + also have "... \ b = a * c" by (simp add: divide_inverse mult.assoc) finally show ?thesis . qed @@ -262,7 +262,7 @@ proof cases assume "a=0" thus ?thesis by simp next - assume "a\0" + assume "a\0" thus ?thesis by (simp add: nonzero_inverse_minus_eq) qed @@ -271,7 +271,7 @@ proof cases assume "a=0" thus ?thesis by simp next - assume "a\0" + assume "a\0" thus ?thesis by (simp add: nonzero_inverse_inverse_eq) qed @@ -394,7 +394,7 @@ lemma divide_minus1 [simp]: "x / - 1 = - x" using nonzero_minus_divide_right [of "1" x] by simp - + end class field_inverse_zero = field + @@ -409,10 +409,10 @@ lemma inverse_mult_distrib [simp]: "inverse (a * b) = inverse a * inverse b" proof cases - assume "a \ 0 & b \ 0" + assume "a \ 0 & b \ 0" thus ?thesis by (simp add: nonzero_inverse_mult_distrib ac_simps) next - assume "~ (a \ 0 & b \ 0)" + assume "~ (a \ 0 & b \ 0)" thus ?thesis by force qed @@ -470,13 +470,13 @@ lemma minus_divide_divide: "(- a) / (- b) = a / b" -apply (cases "b=0", simp) -apply (simp add: nonzero_minus_divide_divide) +apply (cases "b=0", simp) +apply (simp add: nonzero_minus_divide_divide) done lemma inverse_eq_1_iff [simp]: "inverse x = 1 \ x = 1" - by (insert inverse_eq_iff_eq [of x 1], simp) + by (insert inverse_eq_iff_eq [of x 1], simp) lemma divide_eq_0_iff [simp]: "a / b = 0 \ a = 0 \ b = 0" @@ -489,7 +489,7 @@ done lemma divide_cancel_left [simp]: - "c / a = c / b \ c = 0 \ a = b" + "c / a = c / b \ c = 0 \ a = b" apply (cases "c=0", simp) apply (simp add: divide_inverse) done @@ -524,19 +524,19 @@ class linordered_field = field + linordered_idom begin -lemma positive_imp_inverse_positive: - assumes a_gt_0: "0 < a" +lemma positive_imp_inverse_positive: + assumes a_gt_0: "0 < a" shows "0 < inverse a" proof - - have "0 < a * inverse a" + have "0 < a * inverse a" by (simp add: a_gt_0 [THEN less_imp_not_eq2]) - thus "0 < inverse a" + thus "0 < inverse a" by (simp add: a_gt_0 [THEN less_not_sym] zero_less_mult_iff) qed lemma negative_imp_inverse_negative: "a < 0 \ inverse a < 0" - by (insert positive_imp_inverse_positive [of "-a"], + by (insert positive_imp_inverse_positive [of "-a"], simp add: nonzero_inverse_minus_eq less_imp_not_eq) lemma inverse_le_imp_le: @@ -577,7 +577,7 @@ proof fix x::'a have m1: "- (1::'a) < 0" by simp - from add_strict_right_mono[OF m1, where c=x] + from add_strict_right_mono[OF m1, where c=x] have "(- 1) + x < x" by simp thus "\y. y < x" by blast qed @@ -587,7 +587,7 @@ proof fix x::'a have m1: " (1::'a) > 0" by simp - from add_strict_right_mono[OF m1, where c=x] + from add_strict_right_mono[OF m1, where c=x] have "1 + x > x" by simp thus "\y. y > x" by blast qed @@ -606,13 +606,13 @@ lemma inverse_less_imp_less: "inverse a < inverse b \ 0 < a \ b < a" apply (simp add: less_le [of "inverse a"] less_le [of "b"]) -apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq) +apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq) done text{*Both premises are essential. Consider -1 and 1.*} lemma inverse_less_iff_less [simp]: "0 < a \ 0 < b \ inverse a < inverse b \ b < a" - by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) + by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) lemma le_imp_inverse_le: "a \ b \ 0 < a \ inverse b \ inverse a" @@ -620,42 +620,42 @@ lemma inverse_le_iff_le [simp]: "0 < a \ 0 < b \ inverse a \ inverse b \ b \ a" - by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) + by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) text{*These results refer to both operands being negative. The opposite-sign case is trivial, since inverse preserves signs.*} lemma inverse_le_imp_le_neg: "inverse a \ inverse b \ b < 0 \ b \ a" -apply (rule classical) -apply (subgoal_tac "a < 0") +apply (rule classical) +apply (subgoal_tac "a < 0") prefer 2 apply force apply (insert inverse_le_imp_le [of "-b" "-a"]) -apply (simp add: nonzero_inverse_minus_eq) +apply (simp add: nonzero_inverse_minus_eq) done lemma less_imp_inverse_less_neg: "a < b \ b < 0 \ inverse b < inverse a" -apply (subgoal_tac "a < 0") - prefer 2 apply (blast intro: less_trans) +apply (subgoal_tac "a < 0") + prefer 2 apply (blast intro: less_trans) apply (insert less_imp_inverse_less [of "-b" "-a"]) -apply (simp add: nonzero_inverse_minus_eq) +apply (simp add: nonzero_inverse_minus_eq) done lemma inverse_less_imp_less_neg: "inverse a < inverse b \ b < 0 \ b < a" -apply (rule classical) -apply (subgoal_tac "a < 0") +apply (rule classical) +apply (subgoal_tac "a < 0") prefer 2 apply force apply (insert inverse_less_imp_less [of "-b" "-a"]) -apply (simp add: nonzero_inverse_minus_eq) +apply (simp add: nonzero_inverse_minus_eq) done lemma inverse_less_iff_less_neg [simp]: "a < 0 \ b < 0 \ inverse a < inverse b \ b < a" apply (insert inverse_less_iff_less [of "-b" "-a"]) -apply (simp del: inverse_less_iff_less +apply (simp del: inverse_less_iff_less add: nonzero_inverse_minus_eq) done @@ -665,7 +665,7 @@ lemma inverse_le_iff_le_neg [simp]: "a < 0 \ b < 0 \ inverse a \ inverse b \ b \ a" - by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) + by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) lemma one_less_inverse: "0 < a \ a < 1 \ 1 < inverse a" @@ -682,7 +682,7 @@ from assms have "a \ b / c \ a * c \ (b / c) * c" using mult_le_cancel_right [of a c "b * inverse c"] by (auto simp add: field_simps) also have "... \ a * c \ b" - by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) + by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed @@ -693,7 +693,7 @@ from assms have "a < b / c \ a * c < (b / c) * c" using mult_less_cancel_right [of a c "b / c"] by auto also have "... = (a*c < b)" - by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) + by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed @@ -704,7 +704,7 @@ from assms have "a < b / c \ (b / c) * c < a * c" using mult_less_cancel_right [of "b / c" c a] by auto also have "... \ b < a * c" - by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) + by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed @@ -715,7 +715,7 @@ from assms have "a \ b / c \ (b / c) * c \ a * c" using mult_le_cancel_right [of "b * inverse c" c a] by (auto simp add: field_simps) also have "... \ b \ a * c" - by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) + by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed @@ -726,7 +726,7 @@ from assms have "b / c \ a \ (b / c) * c \ a * c" using mult_le_cancel_right [of "b / c" c a] by auto also have "... \ b \ a * c" - by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) + by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed @@ -737,7 +737,7 @@ from assms have "b / c < a \ (b / c) * c < a * c" using mult_less_cancel_right [of "b / c" c a] by auto also have "... \ b < a * c" - by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) + by (simp add: less_imp_not_eq2 [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed @@ -746,9 +746,9 @@ shows "b / c \ a \ a * c \ b" proof - from assms have "b / c \ a \ a * c \ (b / c) * c" - using mult_le_cancel_right [of a c "b / c"] by auto + using mult_le_cancel_right [of a c "b / c"] by auto also have "... \ a * c \ b" - by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) + by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed @@ -759,7 +759,7 @@ from assms have "b / c < a \ a * c < b / c * c" using mult_less_cancel_right [of a c "b / c"] by auto also have "... \ a * c < b" - by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) + by (simp add: less_imp_not_eq [OF assms] divide_inverse mult.assoc) finally show ?thesis . qed @@ -842,7 +842,7 @@ by(simp add:field_simps) lemma divide_nonneg_neg: - "0 <= x ==> y < 0 ==> x / y <= 0" + "0 <= x ==> y < 0 ==> x / y <= 0" by(simp add:field_simps) lemma divide_neg_neg: @@ -855,7 +855,7 @@ lemma divide_strict_right_mono: "[|a < b; 0 < c|] ==> a / c < b / c" -by (simp add: less_imp_not_eq2 divide_inverse mult_strict_right_mono +by (simp add: less_imp_not_eq2 divide_inverse mult_strict_right_mono positive_imp_inverse_positive) @@ -865,7 +865,7 @@ apply (simp add: less_imp_not_eq nonzero_minus_divide_right [symmetric]) done -text{*The last premise ensures that @{term a} and @{term b} +text{*The last premise ensures that @{term a} and @{term b} have the same sign*} lemma divide_strict_left_mono: "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / b" @@ -895,7 +895,7 @@ z < x / y" by(simp add:field_simps) -lemma frac_le: "0 <= x ==> +lemma frac_le: "0 <= x ==> x <= y ==> 0 < w ==> w <= z ==> x / z <= y / w" apply (rule mult_imp_div_pos_le) apply simp @@ -905,7 +905,7 @@ apply simp_all done -lemma frac_less: "0 <= x ==> +lemma frac_less: "0 <= x ==> x < y ==> 0 < w ==> w <= z ==> x / z < y / w" apply (rule mult_imp_div_pos_less) apply simp @@ -915,7 +915,7 @@ apply simp_all done -lemma frac_less2: "0 < x ==> +lemma frac_less2: "0 < x ==> x <= y ==> 0 < w ==> w < z ==> x / z < y / w" apply (rule mult_imp_div_pos_less) apply simp_all @@ -933,7 +933,7 @@ subclass unbounded_dense_linorder proof fix x y :: 'a - from less_add_one show "\y. x < y" .. + from less_add_one show "\y. x < y" .. from less_add_one have "x + (- 1) < (x + 1) + (- 1)" by (rule add_strict_right_mono) then have "x - 1 < x + 1 - 1" by simp then have "x - 1 < x" by (simp add: algebra_simps) @@ -943,14 +943,14 @@ lemma nonzero_abs_inverse: "a \ 0 ==> \inverse a\ = inverse \a\" -apply (auto simp add: neq_iff abs_if nonzero_inverse_minus_eq +apply (auto simp add: neq_iff abs_if nonzero_inverse_minus_eq negative_imp_inverse_negative) -apply (blast intro: positive_imp_inverse_positive elim: less_asym) +apply (blast intro: positive_imp_inverse_positive elim: less_asym) done lemma nonzero_abs_divide: "b \ 0 ==> \a / b\ = \a\ / \b\" - by (simp add: divide_inverse abs_mult nonzero_abs_inverse) + by (simp add: divide_inverse abs_mult nonzero_abs_inverse) lemma field_le_epsilon: assumes e: "\e. 0 < e \ x \ y + e" @@ -1003,10 +1003,10 @@ qed lemma inverse_less_1_iff: "inverse x < 1 \ x \ 0 \ 1 < x" - by (simp add: not_le [symmetric] one_le_inverse_iff) + by (simp add: not_le [symmetric] one_le_inverse_iff) lemma inverse_le_1_iff: "inverse x \ 1 \ x \ 0 \ 1 \ x" - by (simp add: not_less [symmetric] one_less_inverse_iff) + by (simp add: not_less [symmetric] one_less_inverse_iff) lemma [divide_simps]: shows le_divide_eq: "a \ b / c \ (if 0 < c then a * c \ b else if c < 0 then b \ a * c else a \ 0)" @@ -1060,13 +1060,13 @@ "[|a \ b; 0 \ c|] ==> a/c \ b/c" by (force simp add: divide_strict_right_mono le_less) -lemma divide_right_mono_neg: "a <= b +lemma divide_right_mono_neg: "a <= b ==> c <= 0 ==> b / c <= a / c" apply (drule divide_right_mono [of _ _ "- c"]) apply auto done -lemma divide_left_mono_neg: "a <= b +lemma divide_left_mono_neg: "a <= b ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b" apply (drule divide_left_mono [of _ _ "- c"]) apply (auto simp add: mult.commute) @@ -1162,28 +1162,28 @@ by (auto simp add: divide_eq_eq) lemma abs_inverse [simp]: - "\inverse a\ = + "\inverse a\ = inverse \a\" -apply (cases "a=0", simp) -apply (simp add: nonzero_abs_inverse) +apply (cases "a=0", simp) +apply (simp add: nonzero_abs_inverse) done lemma abs_divide [simp]: "\a / b\ = \a\ / \b\" -apply (cases "b=0", simp) -apply (simp add: nonzero_abs_divide) +apply (cases "b=0", simp) +apply (simp add: nonzero_abs_divide) done -lemma abs_div_pos: "0 < y ==> +lemma abs_div_pos: "0 < y ==> \x\ / y = \x / y\" apply (subst abs_divide) apply (simp add: order_less_imp_le) done -lemma zero_le_divide_abs_iff [simp]: "(0 \ a / abs b) = (0 \ a | b = 0)" +lemma zero_le_divide_abs_iff [simp]: "(0 \ a / abs b) = (0 \ a | b = 0)" by (auto simp: zero_le_divide_iff) -lemma divide_le_0_abs_iff [simp]: "(a / abs b \ 0) = (a \ 0 | b = 0)" +lemma divide_le_0_abs_iff [simp]: "(a / abs b \ 0) = (a \ 0 | b = 0)" by (auto simp: divide_le_0_iff) lemma field_le_mult_one_interval: @@ -1208,5 +1208,5 @@ code_identifier code_module Fields \ (SML) Arith and (OCaml) Arith and (Haskell) Arith - + end diff -r 37adca7fd48f -r 651ea265d568 src/HOL/GCD.thy --- a/src/HOL/GCD.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/GCD.thy Tue Mar 10 15:20:40 2015 +0000 @@ -28,7 +28,7 @@ section {* Greatest common divisor and least common multiple *} theory GCD -imports Fact +imports Main begin declare One_nat_def [simp del] @@ -50,7 +50,7 @@ class semiring_gcd = comm_semiring_1 + gcd + assumes gcd_dvd1 [iff]: "gcd a b dvd a" and gcd_dvd2 [iff]: "gcd a b dvd b" - and gcd_greatest: "c dvd a \ c dvd b \ c dvd gcd a b" + and gcd_greatest: "c dvd a \ c dvd b \ c dvd gcd a b" class ring_gcd = comm_ring_1 + semiring_gcd @@ -266,10 +266,10 @@ then show "k dvd gcd m n" by (induct m n rule: gcd_nat_induct) (simp_all add: gcd_non_0_nat dvd_mod gcd_0_nat) qed - + instance int :: ring_gcd by intro_classes (simp_all add: dvd_int_unfold_dvd_nat gcd_int_def gcd_greatest) - + lemma dvd_gcd_D1_nat: "k dvd gcd m n \ (k::nat) dvd m" by (metis gcd_dvd1 dvd_trans) @@ -1753,12 +1753,12 @@ text \Fact aliasses\ - -lemmas gcd_dvd1_nat = gcd_dvd1 [where ?'a = nat] + +lemmas gcd_dvd1_nat = gcd_dvd1 [where ?'a = nat] and gcd_dvd2_nat = gcd_dvd2 [where ?'a = nat] and gcd_greatest_nat = gcd_greatest [where ?'a = nat] -lemmas gcd_dvd1_int = gcd_dvd1 [where ?'a = int] +lemmas gcd_dvd1_int = gcd_dvd1 [where ?'a = int] and gcd_dvd2_int = gcd_dvd2 [where ?'a = int] and gcd_greatest_int = gcd_greatest [where ?'a = int] diff -r 37adca7fd48f -r 651ea265d568 src/HOL/HOLCF/Universal.thy --- a/src/HOL/HOLCF/Universal.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/HOLCF/Universal.thy Tue Mar 10 15:20:40 2015 +0000 @@ -325,8 +325,6 @@ qed qed -no_notation binomial (infixl "choose" 65) - definition choose :: "'a compact_basis set \ 'a compact_basis" where diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Int.thy --- a/src/HOL/Int.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Int.thy Tue Mar 10 15:20:40 2015 +0000 @@ -1599,4 +1599,8 @@ lifting_update int.lifting lifting_forget int.lifting +text{*Also the class for fields with characteristic zero.*} +class field_char_0 = field + ring_char_0 +subclass (in linordered_field) field_char_0 .. + end diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Library/Formal_Power_Series.thy Tue Mar 10 15:20:40 2015 +0000 @@ -5,7 +5,7 @@ section{* A formalization of formal power series *} theory Formal_Power_Series -imports "~~/src/HOL/Number_Theory/Binomial" +imports Complex_Main begin diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Library/NthRoot_Limits.thy --- a/src/HOL/Library/NthRoot_Limits.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Library/NthRoot_Limits.thy Tue Mar 10 15:20:40 2015 +0000 @@ -1,13 +1,7 @@ theory NthRoot_Limits - imports Complex_Main "~~/src/HOL/Number_Theory/Binomial" + imports Complex_Main begin -text {* - -This does not fit into @{text Complex_Main}, as it depends on @{text Binomial} - -*} - lemma LIMSEQ_root: "(\n. root n n) ----> 1" proof - def x \ "\n. root n n - 1" diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Number_Theory/Binomial.thy --- a/src/HOL/Number_Theory/Binomial.thy Tue Mar 10 11:56:32 2015 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,656 +0,0 @@ -(* Title: HOL/Number_Theory/Binomial.thy - Authors: Lawrence C. Paulson, Jeremy Avigad, Tobias Nipkow - -Defines the "choose" function, and establishes basic properties. -*) - -section {* Binomial *} - -theory Binomial -imports Cong Fact Complex_Main -begin - -lemma choose_row_sum: "(\k=0..n. n choose k) = 2^n" - using binomial [of 1 "1" n] - by (simp add: numeral_2_eq_2) - -lemma sum_choose_lower: "(\k=0..n. (r+k) choose k) = Suc (r+n) choose n" - by (induct n) auto - -lemma sum_choose_upper: "(\k=0..n. k choose m) = Suc n choose Suc m" - by (induct n) auto - -lemma natsum_reverse_index: - fixes m::nat - shows "(\k. m \ k \ k \ n \ g k = f (m + n - k)) \ (\k=m..n. f k) = (\k=m..n. g k)" - by (rule setsum.reindex_bij_witness[where i="\k. m+n-k" and j="\k. m+n-k"]) auto - -text{*NW diagonal sum property*} -lemma sum_choose_diagonal: - assumes "m\n" shows "(\k=0..m. (n-k) choose (m-k)) = Suc n choose m" -proof - - have "(\k=0..m. (n-k) choose (m-k)) = (\k=0..m. (n-m+k) choose k)" - by (rule natsum_reverse_index) (simp add: assms) - also have "... = Suc (n-m+m) choose m" - by (rule sum_choose_lower) - also have "... = Suc n choose m" using assms - by simp - finally show ?thesis . -qed - -subsection{* Pochhammer's symbol : generalized rising factorial *} - -text {* See @{url "http://en.wikipedia.org/wiki/Pochhammer_symbol"} *} - -definition "pochhammer (a::'a::comm_semiring_1) n = - (if n = 0 then 1 else setprod (\n. a + of_nat n) {0 .. n - 1})" - -lemma pochhammer_0 [simp]: "pochhammer a 0 = 1" - by (simp add: pochhammer_def) - -lemma pochhammer_1 [simp]: "pochhammer a 1 = a" - by (simp add: pochhammer_def) - -lemma pochhammer_Suc0 [simp]: "pochhammer a (Suc 0) = a" - by (simp add: pochhammer_def) - -lemma pochhammer_Suc_setprod: "pochhammer a (Suc n) = setprod (\n. a + of_nat n) {0 .. n}" - by (simp add: pochhammer_def) - -lemma setprod_nat_ivl_Suc: "setprod f {0 .. Suc n} = setprod f {0..n} * f (Suc n)" -proof - - have "{0..Suc n} = {0..n} \ {Suc n}" by auto - then show ?thesis by (simp add: field_simps) -qed - -lemma setprod_nat_ivl_1_Suc: "setprod f {0 .. Suc n} = f 0 * setprod f {1.. Suc n}" -proof - - have "{0..Suc n} = {0} \ {1 .. Suc n}" by auto - then show ?thesis by simp -qed - - -lemma pochhammer_Suc: "pochhammer a (Suc n) = pochhammer a n * (a + of_nat n)" -proof (cases n) - case 0 - then show ?thesis by simp -next - case (Suc n) - show ?thesis unfolding Suc pochhammer_Suc_setprod setprod_nat_ivl_Suc .. -qed - -lemma pochhammer_rec: "pochhammer a (Suc n) = a * pochhammer (a + 1) n" -proof (cases "n = 0") - case True - then show ?thesis by (simp add: pochhammer_Suc_setprod) -next - case False - have *: "finite {1 .. n}" "0 \ {1 .. n}" by auto - have eq: "insert 0 {1 .. n} = {0..n}" by auto - have **: "(\n\{1\nat..n}. a + of_nat n) = (\n\{0\nat..n - 1}. a + 1 + of_nat n)" - apply (rule setprod.reindex_cong [where l = Suc]) - using False - apply (auto simp add: fun_eq_iff field_simps) - done - show ?thesis - apply (simp add: pochhammer_def) - unfolding setprod.insert [OF *, unfolded eq] - using ** apply (simp add: field_simps) - done -qed - -lemma pochhammer_fact: "of_nat (fact n) = pochhammer 1 n" - unfolding fact_altdef_nat - apply (cases n) - apply (simp_all add: of_nat_setprod pochhammer_Suc_setprod) - apply (rule setprod.reindex_cong [where l = Suc]) - apply (auto simp add: fun_eq_iff) - done - -lemma pochhammer_of_nat_eq_0_lemma: - assumes "k > n" - shows "pochhammer (- (of_nat n :: 'a:: idom)) k = 0" -proof (cases "n = 0") - case True - then show ?thesis - using assms by (cases k) (simp_all add: pochhammer_rec) -next - case False - from assms obtain h where "k = Suc h" by (cases k) auto - then show ?thesis - by (simp add: pochhammer_Suc_setprod) - (metis Suc_leI Suc_le_mono assms atLeastAtMost_iff less_eq_nat.simps(1)) -qed - -lemma pochhammer_of_nat_eq_0_lemma': - assumes kn: "k \ n" - shows "pochhammer (- (of_nat n :: 'a:: {idom,ring_char_0})) k \ 0" -proof (cases k) - case 0 - then show ?thesis by simp -next - case (Suc h) - then show ?thesis - apply (simp add: pochhammer_Suc_setprod) - using Suc kn apply (auto simp add: algebra_simps) - done -qed - -lemma pochhammer_of_nat_eq_0_iff: - shows "pochhammer (- (of_nat n :: 'a:: {idom,ring_char_0})) k = 0 \ k > n" - (is "?l = ?r") - using pochhammer_of_nat_eq_0_lemma[of n k, where ?'a='a] - pochhammer_of_nat_eq_0_lemma'[of k n, where ?'a = 'a] - by (auto simp add: not_le[symmetric]) - - -lemma pochhammer_eq_0_iff: "pochhammer a n = (0::'a::field_char_0) \ (\k < n. a = - of_nat k)" - apply (auto simp add: pochhammer_of_nat_eq_0_iff) - apply (cases n) - apply (auto simp add: pochhammer_def algebra_simps group_add_class.eq_neg_iff_add_eq_0) - apply (metis leD not_less_eq) - done - - -lemma pochhammer_eq_0_mono: - "pochhammer a n = (0::'a::field_char_0) \ m \ n \ pochhammer a m = 0" - unfolding pochhammer_eq_0_iff by auto - -lemma pochhammer_neq_0_mono: - "pochhammer a m \ (0::'a::field_char_0) \ m \ n \ pochhammer a n \ 0" - unfolding pochhammer_eq_0_iff by auto - -lemma pochhammer_minus: - assumes kn: "k \ n" - shows "pochhammer (- b) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (b - of_nat k + 1) k" -proof (cases k) - case 0 - then show ?thesis by simp -next - case (Suc h) - have eq: "((- 1) ^ Suc h :: 'a) = (\i=0..h. - 1)" - using setprod_constant[where A="{0 .. h}" and y="- 1 :: 'a"] - by auto - show ?thesis - unfolding Suc pochhammer_Suc_setprod eq setprod.distrib[symmetric] - by (rule setprod.reindex_bij_witness[where i="op - h" and j="op - h"]) - (auto simp: of_nat_diff) -qed - -lemma pochhammer_minus': - assumes kn: "k \ n" - shows "pochhammer (b - of_nat k + 1) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (- b) k" - unfolding pochhammer_minus[OF kn, where b=b] - unfolding mult.assoc[symmetric] - unfolding power_add[symmetric] - by simp - -lemma pochhammer_same: "pochhammer (- of_nat n) n = - ((- 1) ^ n :: 'a::comm_ring_1) * of_nat (fact n)" - unfolding pochhammer_minus[OF le_refl[of n]] - by (simp add: of_nat_diff pochhammer_fact) - - -subsection{* Generalized binomial coefficients *} - -definition gbinomial :: "'a::field_char_0 \ nat \ 'a" (infixl "gchoose" 65) - where "a gchoose n = - (if n = 0 then 1 else (setprod (\i. a - of_nat i) {0 .. n - 1}) / of_nat (fact n))" - -lemma gbinomial_0 [simp]: "a gchoose 0 = 1" "0 gchoose (Suc n) = 0" - apply (simp_all add: gbinomial_def) - apply (subgoal_tac "(\i\nat\{0\nat..n}. - of_nat i) = (0::'b)") - apply (simp del:setprod_zero_iff) - apply simp - done - -lemma gbinomial_pochhammer: "a gchoose n = (- 1) ^ n * pochhammer (- a) n / of_nat (fact n)" -proof (cases "n = 0") - case True - then show ?thesis by simp -next - case False - from this setprod_constant[of "{0 .. n - 1}" "- (1:: 'a)"] - have eq: "(- (1\'a)) ^ n = setprod (\i. - 1) {0 .. n - 1}" - by auto - from False show ?thesis - by (simp add: pochhammer_def gbinomial_def field_simps - eq setprod.distrib[symmetric]) -qed - - -lemma binomial_gbinomial: "of_nat (n choose k) = of_nat n gchoose k" -proof - - { assume kn: "k > n" - then have ?thesis - by (subst binomial_eq_0[OF kn]) - (simp add: gbinomial_pochhammer field_simps pochhammer_of_nat_eq_0_iff) } - moreover - { assume "k=0" then have ?thesis by simp } - moreover - { assume kn: "k \ n" and k0: "k\ 0" - from k0 obtain h where h: "k = Suc h" by (cases k) auto - from h - have eq:"(- 1 :: 'a) ^ k = setprod (\i. - 1) {0..h}" - by (subst setprod_constant) auto - have eq': "(\i\{0..h}. of_nat n + - (of_nat i :: 'a)) = (\i\{n - h..n}. of_nat i)" - using h kn - by (intro setprod.reindex_bij_witness[where i="op - n" and j="op - n"]) - (auto simp: of_nat_diff) - have th0: "finite {1..n - Suc h}" "finite {n - h .. n}" - "{1..n - Suc h} \ {n - h .. n} = {}" and - eq3: "{1..n - Suc h} \ {n - h .. n} = {1..n}" - using h kn by auto - from eq[symmetric] - have ?thesis using kn - apply (simp add: binomial_fact[OF kn, where ?'a = 'a] - gbinomial_pochhammer field_simps pochhammer_Suc_setprod) - apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h - of_nat_setprod setprod.distrib[symmetric] eq' del: One_nat_def power_Suc) - unfolding setprod.union_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \ 'a"] eq[unfolded h] - unfolding mult.assoc[symmetric] - unfolding setprod.distrib[symmetric] - apply simp - apply (intro setprod.reindex_bij_witness[where i="op - n" and j="op - n"]) - apply (auto simp: of_nat_diff) - done - } - moreover - have "k > n \ k = 0 \ (k \ n \ k \ 0)" by arith - ultimately show ?thesis by blast -qed - -lemma gbinomial_1[simp]: "a gchoose 1 = a" - by (simp add: gbinomial_def) - -lemma gbinomial_Suc0[simp]: "a gchoose (Suc 0) = a" - by (simp add: gbinomial_def) - -lemma gbinomial_mult_1: - "a * (a gchoose n) = - of_nat n * (a gchoose n) + of_nat (Suc n) * (a gchoose (Suc n))" (is "?l = ?r") -proof - - have "?r = ((- 1) ^n * pochhammer (- a) n / of_nat (fact n)) * (of_nat n - (- a + of_nat n))" - unfolding gbinomial_pochhammer - pochhammer_Suc fact_Suc of_nat_mult right_diff_distrib power_Suc - by (simp add: field_simps del: of_nat_Suc) - also have "\ = ?l" unfolding gbinomial_pochhammer - by (simp add: field_simps) - finally show ?thesis .. -qed - -lemma gbinomial_mult_1': - "(a gchoose n) * a = of_nat n * (a gchoose n) + of_nat (Suc n) * (a gchoose (Suc n))" - by (simp add: mult.commute gbinomial_mult_1) - -lemma gbinomial_Suc: - "a gchoose (Suc k) = (setprod (\i. a - of_nat i) {0 .. k}) / of_nat (fact (Suc k))" - by (simp add: gbinomial_def) - -lemma gbinomial_mult_fact: - "(of_nat (fact (Suc k)) :: 'a) * ((a::'a::field_char_0) gchoose (Suc k)) = - (setprod (\i. a - of_nat i) {0 .. k})" - by (simp_all add: gbinomial_Suc field_simps del: fact_Suc) - -lemma gbinomial_mult_fact': - "((a::'a::field_char_0) gchoose (Suc k)) * (of_nat (fact (Suc k)) :: 'a) = - (setprod (\i. a - of_nat i) {0 .. k})" - using gbinomial_mult_fact[of k a] - by (subst mult.commute) - - -lemma gbinomial_Suc_Suc: - "((a::'a::field_char_0) + 1) gchoose (Suc k) = a gchoose k + (a gchoose (Suc k))" -proof (cases k) - case 0 - then show ?thesis by simp -next - case (Suc h) - have eq0: "(\i\{1..k}. (a + 1) - of_nat i) = (\i\{0..h}. a - of_nat i)" - apply (rule setprod.reindex_cong [where l = Suc]) - using Suc - apply auto - done - have "of_nat (fact (Suc k)) * (a gchoose k + (a gchoose (Suc k))) = - ((a gchoose Suc h) * of_nat (fact (Suc h)) * of_nat (Suc k)) + (\i\{0\nat..Suc h}. a - of_nat i)" - apply (simp add: Suc field_simps del: fact_Suc) - unfolding gbinomial_mult_fact' - apply (subst fact_Suc) - unfolding of_nat_mult - apply (subst mult.commute) - unfolding mult.assoc - unfolding gbinomial_mult_fact - apply (simp add: field_simps) - done - also have "\ = (\i\{0..h}. a - of_nat i) * (a + 1)" - unfolding gbinomial_mult_fact' setprod_nat_ivl_Suc - by (simp add: field_simps Suc) - also have "\ = (\i\{0..k}. (a + 1) - of_nat i)" - using eq0 - by (simp add: Suc setprod_nat_ivl_1_Suc) - also have "\ = of_nat (fact (Suc k)) * ((a + 1) gchoose (Suc k))" - unfolding gbinomial_mult_fact .. - finally show ?thesis by (simp del: fact_Suc) -qed - -lemma gbinomial_reduce_nat: - "0 < k \ (a::'a::field_char_0) gchoose k = (a - 1) gchoose (k - 1) + ((a - 1) gchoose k)" -by (metis Suc_pred' diff_add_cancel gbinomial_Suc_Suc) - - -lemma binomial_symmetric: - assumes kn: "k \ n" - shows "n choose k = n choose (n - k)" -proof- - from kn have kn': "n - k \ n" by arith - from binomial_fact_lemma[OF kn] binomial_fact_lemma[OF kn'] - have "fact k * fact (n - k) * (n choose k) = - fact (n - k) * fact (n - (n - k)) * (n choose (n - k))" by simp - then show ?thesis using kn by simp -qed - -text{*Contributed by Manuel Eberl, generalised by LCP. - Alternative definition of the binomial coefficient as @{term "\iii = (\i x" - shows "(x / of_nat k :: 'a) ^ k \ x gchoose k" -proof - - have x: "0 \ x" - using assms of_nat_0_le_iff order_trans by blast - have "(x / of_nat k :: 'a) ^ k = (\i \ x gchoose k" - unfolding gbinomial_altdef_of_nat - proof (safe intro!: setprod_mono) - fix i :: nat - assume ik: "i < k" - from assms have "x * of_nat i \ of_nat (i * k)" - by (metis mult.commute mult_le_cancel_right of_nat_less_0_iff of_nat_mult) - then have "x * of_nat k - x * of_nat i \ x * of_nat k - of_nat (i * k)" by arith - then have "x * of_nat (k - i) \ (x - of_nat i) * of_nat k" - using ik - by (simp add: algebra_simps zero_less_mult_iff of_nat_diff of_nat_mult) - then have "x * of_nat (k - i) \ (x - of_nat i) * (of_nat k :: 'a)" - unfolding of_nat_mult[symmetric] of_nat_le_iff . - with assms show "x / of_nat k \ (x - of_nat i) / (of_nat (k - i) :: 'a)" - using `i < k` by (simp add: field_simps) - qed (simp add: x zero_le_divide_iff) - finally show ?thesis . -qed - -text{*Versions of the theorems above for the natural-number version of "choose"*} -lemma binomial_altdef_of_nat: - fixes n k :: nat - and x :: "'a :: {field_char_0,field_inverse_zero}" --{*the point is to constrain @{typ 'a}*} - assumes "k \ n" - shows "of_nat (n choose k) = (\i n" - shows "(of_nat n / of_nat k :: 'a) ^ k \ of_nat (n choose k)" -by (simp add: assms gbinomial_ge_n_over_k_pow_k binomial_gbinomial of_nat_diff) - -lemma binomial_le_pow: - assumes "r \ n" - shows "n choose r \ n ^ r" -proof - - have "n choose r \ fact n div fact (n - r)" - using `r \ n` by (subst binomial_fact_lemma[symmetric]) auto - with fact_div_fact_le_pow [OF assms] show ?thesis by auto -qed - -lemma binomial_altdef_nat: "(k::nat) \ n \ - n choose k = fact n div (fact k * fact (n - k))" - by (subst binomial_fact_lemma [symmetric]) auto - -lemma choose_dvd_nat: "(k::nat) \ n \ fact k * fact (n - k) dvd fact n" -by (metis binomial_fact_lemma dvd_def) - -lemma choose_dvd_int: - assumes "(0::int) <= k" and "k <= n" - shows "fact k * fact (n - k) dvd fact n" - apply (subst tsub_eq [symmetric], rule assms) - apply (rule choose_dvd_nat [transferred]) - using assms apply auto - done - -lemma fact_fact_dvd_fact: fixes k::nat shows "fact k * fact n dvd fact (n + k)" -by (metis add.commute add_diff_cancel_left' choose_dvd_nat le_add2) - -lemma choose_mult_lemma: - "((m+r+k) choose (m+k)) * ((m+k) choose k) = ((m+r+k) choose k) * ((m+r) choose m)" -proof - - have "((m+r+k) choose (m+k)) * ((m+k) choose k) = - fact (m+r + k) div (fact (m + k) * fact (m+r - m)) * (fact (m + k) div (fact k * fact m))" - by (simp add: assms binomial_altdef_nat) - also have "... = fact (m+r+k) div (fact r * (fact k * fact m))" - apply (subst div_mult_div_if_dvd) - apply (auto simp: fact_fact_dvd_fact) - apply (metis add.assoc add.commute fact_fact_dvd_fact) - done - also have "... = (fact (m+r+k) * fact (m+r)) div (fact r * (fact k * fact m) * fact (m+r))" - apply (subst div_mult_div_if_dvd [symmetric]) - apply (auto simp: fact_fact_dvd_fact) - apply (metis dvd_trans dvd.dual_order.refl fact_fact_dvd_fact mult_dvd_mono mult.left_commute) - done - also have "... = (fact (m+r+k) div (fact k * fact (m+r)) * (fact (m+r) div (fact r * fact m)))" - apply (subst div_mult_div_if_dvd) - apply (auto simp: fact_fact_dvd_fact) - apply(metis mult.left_commute) - done - finally show ?thesis - by (simp add: binomial_altdef_nat mult.commute) -qed - -text{*The "Subset of a Subset" identity*} -lemma choose_mult: - assumes "k\m" "m\n" - shows "(n choose m) * (m choose k) = (n choose k) * ((n-k) choose (m-k))" -using assms choose_mult_lemma [of "m-k" "n-m" k] -by simp - - -subsection {* Binomial coefficients *} - -lemma choose_one: "(n::nat) choose 1 = n" - by simp - -(*FIXME: messy and apparently unused*) -lemma binomial_induct [rule_format]: "(ALL (n::nat). P n n) \ - (ALL n. P (Suc n) 0) \ (ALL n. (ALL k < n. P n k \ P n (Suc k) \ - P (Suc n) (Suc k))) \ (ALL k <= n. P n k)" - apply (induct n) - apply auto - apply (case_tac "k = 0") - apply auto - apply (case_tac "k = Suc n") - apply auto - apply (metis Suc_le_eq fact_nat.cases le_Suc_eq le_eq_less_or_eq) - done - -lemma card_UNION: - assumes "finite A" and "\k \ A. finite k" - shows "card (\A) = nat (\I | I \ A \ I \ {}. (- 1) ^ (card I + 1) * int (card (\I)))" - (is "?lhs = ?rhs") -proof - - have "?rhs = nat (\I | I \ A \ I \ {}. (- 1) ^ (card I + 1) * (\_\\I. 1))" by simp - also have "\ = nat (\I | I \ A \ I \ {}. (\_\\I. (- 1) ^ (card I + 1)))" (is "_ = nat ?rhs") - by(subst setsum_right_distrib) simp - also have "?rhs = (\(I, _)\Sigma {I. I \ A \ I \ {}} Inter. (- 1) ^ (card I + 1))" - using assms by(subst setsum.Sigma)(auto) - also have "\ = (\(x, I)\(SIGMA x:UNIV. {I. I \ A \ I \ {} \ x \ \I}). (- 1) ^ (card I + 1))" - by (rule setsum.reindex_cong [where l = "\(x, y). (y, x)"]) (auto intro: inj_onI simp add: split_beta) - also have "\ = (\(x, I)\(SIGMA x:\A. {I. I \ A \ I \ {} \ x \ \I}). (- 1) ^ (card I + 1))" - using assms by(auto intro!: setsum.mono_neutral_cong_right finite_SigmaI2 intro: finite_subset[where B="\A"]) - also have "\ = (\x\\A. (\I|I \ A \ I \ {} \ x \ \I. (- 1) ^ (card I + 1)))" - using assms by(subst setsum.Sigma) auto - also have "\ = (\_\\A. 1)" (is "setsum ?lhs _ = _") - proof(rule setsum.cong[OF refl]) - fix x - assume x: "x \ \A" - def K \ "{X \ A. x \ X}" - with `finite A` have K: "finite K" by auto - let ?I = "\i. {I. I \ A \ card I = i \ x \ \I}" - have "inj_on snd (SIGMA i:{1..card A}. ?I i)" - using assms by(auto intro!: inj_onI) - moreover have [symmetric]: "snd ` (SIGMA i:{1..card A}. ?I i) = {I. I \ A \ I \ {} \ x \ \I}" - using assms by(auto intro!: rev_image_eqI[where x="(card a, a)" for a] - simp add: card_gt_0_iff[folded Suc_le_eq] - dest: finite_subset intro: card_mono) - ultimately have "?lhs x = (\(i, I)\(SIGMA i:{1..card A}. ?I i). (- 1) ^ (i + 1))" - by (rule setsum.reindex_cong [where l = snd]) fastforce - also have "\ = (\i=1..card A. (\I|I \ A \ card I = i \ x \ \I. (- 1) ^ (i + 1)))" - using assms by(subst setsum.Sigma) auto - also have "\ = (\i=1..card A. (- 1) ^ (i + 1) * (\I|I \ A \ card I = i \ x \ \I. 1))" - by(subst setsum_right_distrib) simp - also have "\ = (\i=1..card K. (- 1) ^ (i + 1) * (\I|I \ K \ card I = i. 1))" (is "_ = ?rhs") - proof(rule setsum.mono_neutral_cong_right[rule_format]) - show "{1..card K} \ {1..card A}" using `finite A` - by(auto simp add: K_def intro: card_mono) - next - fix i - assume "i \ {1..card A} - {1..card K}" - hence i: "i \ card A" "card K < i" by auto - have "{I. I \ A \ card I = i \ x \ \I} = {I. I \ K \ card I = i}" - by(auto simp add: K_def) - also have "\ = {}" using `finite A` i - by(auto simp add: K_def dest: card_mono[rotated 1]) - finally show "(- 1) ^ (i + 1) * (\I | I \ A \ card I = i \ x \ \I. 1 :: int) = 0" - by(simp only:) simp - next - fix i - have "(\I | I \ A \ card I = i \ x \ \I. 1) = (\I | I \ K \ card I = i. 1 :: int)" - (is "?lhs = ?rhs") - by(rule setsum.cong)(auto simp add: K_def) - thus "(- 1) ^ (i + 1) * ?lhs = (- 1) ^ (i + 1) * ?rhs" by simp - qed simp - also have "{I. I \ K \ card I = 0} = {{}}" using assms - by(auto simp add: card_eq_0_iff K_def dest: finite_subset) - hence "?rhs = (\i = 0..card K. (- 1) ^ (i + 1) * (\I | I \ K \ card I = i. 1 :: int)) + 1" - by(subst (2) setsum_head_Suc)(simp_all ) - also have "\ = (\i = 0..card K. (- 1) * ((- 1) ^ i * int (card K choose i))) + 1" - using K by(subst n_subsets[symmetric]) simp_all - also have "\ = - (\i = 0..card K. (- 1) ^ i * int (card K choose i)) + 1" - by(subst setsum_right_distrib[symmetric]) simp - also have "\ = - ((-1 + 1) ^ card K) + 1" - by(subst binomial_ring)(simp add: ac_simps) - also have "\ = 1" using x K by(auto simp add: K_def card_gt_0_iff) - finally show "?lhs x = 1" . - qed - also have "nat \ = card (\A)" by simp - finally show ?thesis .. -qed - -text{* The number of nat lists of length @{text m} summing to @{text N} is -@{term "(N + m - 1) choose N"}: *} - -lemma card_length_listsum_rec: - assumes "m\1" - shows "card {l::nat list. length l = m \ listsum l = N} = - (card {l. length l = (m - 1) \ listsum l = N} + - card {l. length l = m \ listsum l + 1 = N})" - (is "card ?C = (card ?A + card ?B)") -proof - - let ?A'="{l. length l = m \ listsum l = N \ hd l = 0}" - let ?B'="{l. length l = m \ listsum l = N \ hd l \ 0}" - let ?f ="\ l. 0#l" - let ?g ="\ l. (hd l + 1) # tl l" - have 1: "\xs x. xs \ [] \ x = hd xs \ x # tl xs = xs" by simp - have 2: "\xs. (xs::nat list) \ [] \ listsum(tl xs) = listsum xs - hd xs" - by(auto simp add: neq_Nil_conv) - have f: "bij_betw ?f ?A ?A'" - apply(rule bij_betw_byWitness[where f' = tl]) - using assms - by (auto simp: 2 length_0_conv[symmetric] 1 simp del: length_0_conv) - have 3: "\xs:: nat list. xs \ [] \ hd xs + (listsum xs - hd xs) = listsum xs" - by (metis 1 listsum_simps(2) 2) - have g: "bij_betw ?g ?B ?B'" - apply(rule bij_betw_byWitness[where f' = "\ l. (hd l - 1) # tl l"]) - using assms - by (auto simp: 2 length_0_conv[symmetric] intro!: 3 - simp del: length_greater_0_conv length_0_conv) - { fix M N :: nat have "finite {xs. size xs = M \ set xs \ {0.. ?B'" by auto - have disj: "?A' \ ?B' = {}" by auto - have "card ?C = card(?A' \ ?B')" using uni by simp - also have "\ = card ?A + card ?B" - using card_Un_disjoint[OF _ _ disj] bij_betw_finite[OF f] bij_betw_finite[OF g] - bij_betw_same_card[OF f] bij_betw_same_card[OF g] fin_A fin_B - by presburger - finally show ?thesis . -qed - -lemma card_length_listsum: --"By Holden Lee, tidied by Tobias Nipkow" - "card {l::nat list. size l = m \ listsum l = N} = (N + m - 1) choose N" -proof (cases m) - case 0 then show ?thesis - by (cases N) (auto simp: cong: conj_cong) -next - case (Suc m') - have m: "m\1" by (simp add: Suc) - then show ?thesis - proof (induct "N + m - 1" arbitrary: N m) - case 0 -- "In the base case, the only solution is [0]." - have [simp]: "{l::nat list. length l = Suc 0 \ (\n\set l. n = 0)} = {[0]}" - by (auto simp: length_Suc_conv) - have "m=1 \ N=0" using 0 by linarith - then show ?case by simp - next - case (Suc k) - - have c1: "card {l::nat list. size l = (m - 1) \ listsum l = N} = - (N + (m - 1) - 1) choose N" - proof cases - assume "m = 1" - with Suc.hyps have "N\1" by auto - with `m = 1` show ?thesis by (simp add: binomial_eq_0) - next - assume "m \ 1" thus ?thesis using Suc by fastforce - qed - - from Suc have c2: "card {l::nat list. size l = m \ listsum l + 1 = N} = - (if N>0 then ((N - 1) + m - 1) choose (N - 1) else 0)" - proof - - have aux: "\m n. n > 0 \ Suc m = n \ m = n - 1" by arith - from Suc have "N>0 \ - card {l::nat list. size l = m \ listsum l + 1 = N} = - ((N - 1) + m - 1) choose (N - 1)" by (simp add: aux) - thus ?thesis by auto - qed - - from Suc.prems have "(card {l::nat list. size l = (m - 1) \ listsum l = N} + - card {l::nat list. size l = m \ listsum l + 1 = N}) = (N + m - 1) choose N" - by (auto simp: c1 c2 choose_reduce_nat[of "N + m - 1" N] simp del: One_nat_def) - thus ?case using card_length_listsum_rec[OF Suc.prems] by auto - qed -qed - -end diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Number_Theory/Cong.thy --- a/src/HOL/Number_Theory/Cong.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Number_Theory/Cong.thy Tue Mar 10 15:20:40 2015 +0000 @@ -461,15 +461,7 @@ lemma cong_to_1'_nat: "[(a::nat) = 1] (mod n) \ a = 0 \ n = 1 \ (\m. a = 1 + m * n)" - apply (cases "n = 1") - apply auto [1] - apply (drule_tac x = "a - 1" in spec) - apply force - apply (cases "a = 0", simp add: cong_0_1_nat) - apply (rule iffI) - apply (metis cong_to_1_nat dvd_def monoid_mult_class.mult.right_neutral mult.commute mult_eq_if) - apply (metis cong_add_lcancel_0_nat cong_mult_self_nat) - done +by (metis add.right_neutral cong_0_1_nat cong_iff_lin_nat cong_to_1_nat dvd_div_mult_self leI le_add_diff_inverse less_one mult_eq_if) lemma cong_le_nat: "(y::nat) <= x \ [x = y] (mod n) \ (\q. x = q * n + y)" by (metis cong_altdef_nat Nat.le_imp_diff_is_add dvd_def mult.commute) @@ -579,7 +571,7 @@ [x = y] (mod (PROD i:A. m i))" apply (induct set: finite) apply auto - apply (metis coprime_cong_mult_nat gcd_semilattice_nat.inf_commute setprod_coprime_nat) + apply (metis One_nat_def coprime_cong_mult_nat gcd_nat.commute setprod_coprime_nat) done lemma cong_cong_setprod_coprime_int [rule_format]: "finite A \ @@ -835,7 +827,7 @@ [x = y] (mod (PROD i:A. m i))" apply (induct set: finite) apply auto - apply (metis coprime_cong_mult_nat mult.commute setprod_coprime_nat) + apply (metis One_nat_def coprime_cong_mult_nat gcd_nat.commute setprod_coprime_nat) done lemma chinese_remainder_unique_nat: diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Number_Theory/Fib.thy --- a/src/HOL/Number_Theory/Fib.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Number_Theory/Fib.thy Tue Mar 10 15:20:40 2015 +0000 @@ -11,7 +11,7 @@ section {* Fib *} theory Fib -imports Binomial +imports Main "../GCD" begin diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Number_Theory/Primes.thy --- a/src/HOL/Number_Theory/Primes.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Number_Theory/Primes.thy Tue Mar 10 15:20:40 2015 +0000 @@ -28,7 +28,7 @@ section {* Primes *} theory Primes -imports "~~/src/HOL/GCD" +imports "~~/src/HOL/GCD" "~~/src/HOL/Fact" begin declare [[coercion int]] diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Number_Theory/Residues.thy --- a/src/HOL/Number_Theory/Residues.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Number_Theory/Residues.thy Tue Mar 10 15:20:40 2015 +0000 @@ -8,10 +8,7 @@ section {* Residue rings *} theory Residues -imports - UniqueFactorization - Binomial - MiscAlgebra +imports UniqueFactorization MiscAlgebra begin (* @@ -275,15 +272,15 @@ then have cop: "\x. x \ {1::nat..p - 1} \ coprime x p" by blast { fix x::nat assume *: "1 < x" "x < p" and "x dvd p" - have "coprime x p" + have "coprime x p" apply (rule cop) using * apply auto done with `x dvd p` `1 < x` have "False" by auto } - then show ?thesis - using `2 \ p` + then show ?thesis + using `2 \ p` by (simp add: prime_def) - (metis One_nat_def dvd_pos_nat nat_dvd_not_less nat_neq_iff not_gr0 + (metis One_nat_def dvd_pos_nat nat_dvd_not_less nat_neq_iff not_gr0 not_numeral_le_zero one_dvd) qed @@ -367,7 +364,7 @@ also have "phi p = nat p - 1" by (rule phi_prime, rule assms) finally show ?thesis - by (metis nat_int) + by (metis nat_int) qed lemma fermat_theorem_nat: @@ -441,7 +438,7 @@ lemma wilson_theorem: assumes "prime p" shows "[fact (p - 1) = - 1] (mod p)" proof (cases "p = 2") - case True + case True then show ?thesis by (simp add: cong_int_def fact_altdef_nat) next diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Tue Mar 10 15:20:40 2015 +0000 @@ -1,5 +1,5 @@ (* Title: HOL/Probability/Probability_Mass_Function.thy - Author: Johannes Hölzl, TU München + Author: Johannes Hölzl, TU München Author: Andreas Lochbihler, ETH Zurich *) @@ -8,7 +8,6 @@ theory Probability_Mass_Function imports Giry_Monad - "~~/src/HOL/Number_Theory/Binomial" "~~/src/HOL/Library/Multiset" begin @@ -52,14 +51,14 @@ fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X") then obtain X where "finite X" "card X = Suc (Suc n)" "X \ ?X" by (metis infinite_arbitrarily_large) - from this(3) have *: "\x. x \ X \ ?M / Suc n \ ?m x" + from this(3) have *: "\x. x \ X \ ?M / Suc n \ ?m x" by auto { fix x assume "x \ X" from `?M \ 0` *[OF this] have "?m x \ 0" by (auto simp: field_simps measure_le_0_iff) then have "{x} \ sets M" by (auto dest: measure_notin_sets) } note singleton_sets = this have "?M < (\x\X. ?M / Suc n)" - using `?M \ 0` + using `?M \ 0` by (simp add: `card X = Suc (Suc n)` real_eq_of_nat[symmetric] real_of_nat_Suc field_simps less_le measure_nonneg) also have "\ \ (\x\X. ?m x)" by (rule setsum_mono) fact @@ -82,7 +81,7 @@ assume "\S. countable S \ (AE x in M. x \ S)" then obtain S where S[intro]: "countable S" and ae: "AE x in M. x \ S" by auto - then have "emeasure M (\x\{x\S. emeasure M {x} \ 0}. {x}) = + then have "emeasure M (\x\{x\S. emeasure M {x} \ 0}. {x}) = (\\<^sup>+ x. emeasure M {x} * indicator {x\S. emeasure M {x} \ 0} x \count_space UNIV)" by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def nn_integral_restrict_space[symmetric] restrict_count_space) @@ -136,7 +135,7 @@ interpretation pmf_as_measure . lemma sets_measure_pmf[simp]: "sets (measure_pmf p) = UNIV" - by transfer blast + by transfer blast lemma sets_measure_pmf_count_space[measurable_cong]: "sets (measure_pmf M) = sets (count_space UNIV)" @@ -353,10 +352,10 @@ have [measurable]: "g \ measurable f (subprob_algebra (count_space UNIV))" by (auto simp: measurable_def space_subprob_algebra prob_space_imp_subprob_space g) - + show "prob_space (f \= g)" using g by (intro f.prob_space_bind[where S="count_space UNIV"]) auto - then interpret fg: prob_space "f \= g" . + then interpret fg: prob_space "f \= g" . show [simp]: "sets (f \= g) = UNIV" using sets_eq_imp_space_eq[OF s_f] by (subst sets_bind[where N="count_space UNIV"]) auto @@ -385,7 +384,7 @@ by transfer (simp add: bind_const' prob_space_imp_subprob_space) lemma set_bind_pmf[simp]: "set_pmf (bind_pmf M N) = (\M\set_pmf M. set_pmf (N M))" - unfolding set_pmf_eq ereal_eq_0(1)[symmetric] ereal_pmf_bind + unfolding set_pmf_eq ereal_eq_0(1)[symmetric] ereal_pmf_bind by (auto simp add: nn_integral_0_iff_AE AE_measure_pmf_iff set_pmf_eq not_le less_le pmf_nonneg) lemma bind_pmf_cong: @@ -415,7 +414,7 @@ using measurable_measure_pmf[of N] unfolding measure_pmf_bind by (subst emeasure_bind[where N="count_space UNIV"]) auto - + lift_definition return_pmf :: "'a \ 'a pmf" is "return (count_space UNIV)" by (auto intro!: prob_space_return simp: AE_return measure_return) @@ -451,7 +450,7 @@ proof - have "rel_fun op = (rel_fun pmf_as_measure.cr_pmf pmf_as_measure.cr_pmf) (\f M. M \= (return (count_space UNIV) o f)) map_pmf" - unfolding map_pmf_def[abs_def] comp_def by transfer_prover + unfolding map_pmf_def[abs_def] comp_def by transfer_prover then show ?thesis by (force simp: rel_fun_def cr_pmf_def bind_return_distr) qed @@ -468,7 +467,7 @@ using map_pmf_id unfolding id_def . lemma map_pmf_compose: "map_pmf (f \ g) = map_pmf f \ map_pmf g" - by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def) + by (rule, transfer) (simp add: distr_distr[symmetric, where N="count_space UNIV"] measurable_def) lemma map_pmf_comp: "map_pmf f (map_pmf g M) = map_pmf (\x. f (g x)) M" using map_pmf_compose[of f g] by (simp add: comp_def) @@ -665,7 +664,7 @@ show "M = density (count_space UNIV) (\x. ereal (measure M {x}))" proof (rule measure_eqI) fix A :: "'a set" - have "(\\<^sup>+ x. ereal (measure M {x}) * indicator A x \count_space UNIV) = + have "(\\<^sup>+ x. ereal (measure M {x}) * indicator A x \count_space UNIV) = (\\<^sup>+ x. emeasure M {x} * indicator (A \ {x. measure M {x} \ 0}) x \count_space UNIV)" by (auto intro!: nn_integral_cong simp: emeasure_eq_measure split: split_indicator) also have "\ = (\\<^sup>+ x. emeasure M {x} \count_space (A \ {x. measure M {x} \ 0}))" @@ -706,9 +705,9 @@ setup_lifting td_pmf_embed_pmf -lemma set_pmf_transfer[transfer_rule]: +lemma set_pmf_transfer[transfer_rule]: assumes "bi_total A" - shows "rel_fun (pcr_pmf A) (rel_set A) (\f. {x. f x \ 0}) set_pmf" + shows "rel_fun (pcr_pmf A) (rel_set A) (\f. {x. f x \ 0}) set_pmf" using `bi_total A` by (auto simp: pcr_pmf_def cr_pmf_def rel_fun_def rel_set_def bi_total_def Bex_def set_pmf_iff) metis+ @@ -888,14 +887,14 @@ inductive rel_pmf :: "('a \ 'b \ bool) \ 'a pmf \ 'b pmf \ bool" for R p q where - "\ \x y. (x, y) \ set_pmf pq \ R x y; + "\ \x y. (x, y) \ set_pmf pq \ R x y; map_pmf fst pq = p; map_pmf snd pq = q \ \ rel_pmf R p q" bnf pmf: "'a pmf" map: map_pmf sets: set_pmf bd : "natLeq" rel: rel_pmf proof - show "map_pmf id = id" by (rule map_pmf_id) - show "\f g. map_pmf (f \ g) = map_pmf f \ map_pmf g" by (rule map_pmf_compose) + show "\f g. map_pmf (f \ g) = map_pmf f \ map_pmf g" by (rule map_pmf_compose) show "\f g::'a \ 'b. \p. (\x. x \ set_pmf p \ f x = g x) \ map_pmf f p = map_pmf g p" by (intro map_pmf_cong refl) @@ -1042,7 +1041,7 @@ map_pair) qed -lemma rel_pmf_reflI: +lemma rel_pmf_reflI: assumes "\x. x \ set_pmf p \ P x x" shows "rel_pmf P p p" by (rule rel_pmf.intros[where pq="map_pmf (\x. (x, x)) p"]) @@ -1089,7 +1088,7 @@ and q: "q = map_pmf snd pq" and P: "\x y. (x, y) \ set_pmf pq \ rel_pmf P x y" by cases auto - from P obtain PQ + from P obtain PQ where PQ: "\x y a b. \ (x, y) \ set_pmf pq; (a, b) \ set_pmf (PQ x y) \ \ P a b" and x: "\x y. (x, y) \ set_pmf pq \ map_pmf fst (PQ x y) = x" and y: "\x y. (x, y) \ set_pmf pq \ map_pmf snd (PQ x y) = y" @@ -1112,12 +1111,12 @@ text {* Proof that @{const rel_pmf} preserves orders. - Antisymmetry proof follows Thm. 1 in N. Saheb-Djahromi, Cpo's of measures for nondeterminism, - Theoretical Computer Science 12(1):19--37, 1980, + Antisymmetry proof follows Thm. 1 in N. Saheb-Djahromi, Cpo's of measures for nondeterminism, + Theoretical Computer Science 12(1):19--37, 1980, @{url "http://dx.doi.org/10.1016/0304-3975(80)90003-1"} *} -lemma +lemma assumes *: "rel_pmf R p q" and refl: "reflp R" and trans: "transp R" shows measure_Ici: "measure p {y. R x y} \ measure q {y. R x y}" (is ?thesis1) @@ -1174,7 +1173,7 @@ hence "measure (measure_pmf p) (?E x) \ 0" by (auto simp add: measure_pmf.prob_eq_0 AE_measure_pmf_iff intro: reflpD[OF \reflp R\]) hence "measure (measure_pmf q) (?E x) \ 0" using eq by simp - hence "set_pmf q \ {y. R x y \ R y x} \ {}" + hence "set_pmf q \ {y. R x y \ R y x} \ {}" by (auto simp add: measure_pmf.prob_eq_0 AE_measure_pmf_iff) } ultimately show "inf R R\\ x y" by (auto simp add: pq_def) @@ -1235,13 +1234,13 @@ lemma set_pmf_bernoulli: "0 < p \ p < 1 \ set_pmf (bernoulli_pmf p) = UNIV" by (auto simp add: set_pmf_iff UNIV_bool) -lemma nn_integral_bernoulli_pmf[simp]: +lemma nn_integral_bernoulli_pmf[simp]: assumes [simp]: "0 \ p" "p \ 1" "\x. 0 \ f x" shows "(\\<^sup>+x. f x \bernoulli_pmf p) = f True * p + f False * (1 - p)" by (subst nn_integral_measure_pmf_support[of UNIV]) (auto simp: UNIV_bool field_simps) -lemma integral_bernoulli_pmf[simp]: +lemma integral_bernoulli_pmf[simp]: assumes [simp]: "0 \ p" "p \ 1" shows "(\x. f x \bernoulli_pmf p) = f True * p + f False * (1 - p)" by (subst integral_measure_pmf[of UNIV]) (auto simp: UNIV_bool) @@ -1277,7 +1276,7 @@ lift_definition pmf_of_multiset :: "'a pmf" is "\x. count M x / size M" proof - show "(\\<^sup>+ x. ereal (real (count M x) / real (size M)) \count_space UNIV) = 1" + show "(\\<^sup>+ x. ereal (real (count M x) / real (size M)) \count_space UNIV) = 1" using M_not_empty by (simp add: zero_less_divide_iff nn_integral_count_space nonempty_has_size setsum_divide_distrib[symmetric]) @@ -1300,7 +1299,7 @@ lift_definition pmf_of_set :: "'a pmf" is "\x. indicator S x / card S" proof - show "(\\<^sup>+ x. ereal (indicator S x / real (card S)) \count_space UNIV) = 1" + show "(\\<^sup>+ x. ereal (indicator S x / real (card S)) \count_space UNIV) = 1" using S_not_empty S_finite by (subst nn_integral_count_space'[of S]) auto qed simp diff -r 37adca7fd48f -r 651ea265d568 src/HOL/ROOT --- a/src/HOL/ROOT Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/ROOT Tue Mar 10 15:20:40 2015 +0000 @@ -285,7 +285,6 @@ (* Preliminaries from set and number theory *) "~~/src/HOL/Library/FuncSet" "~~/src/HOL/Number_Theory/Primes" - "~~/src/HOL/Number_Theory/Binomial" "~~/src/HOL/Library/Permutation" theories (*** New development, based on explicit structures ***) diff -r 37adca7fd48f -r 651ea265d568 src/HOL/Rat.thy --- a/src/HOL/Rat.thy Tue Mar 10 11:56:32 2015 +0100 +++ b/src/HOL/Rat.thy Tue Mar 10 15:20:40 2015 +0000 @@ -638,10 +638,6 @@ subsection {* Embedding from Rationals to other Fields *} -class field_char_0 = field + ring_char_0 - -subclass (in linordered_field) field_char_0 .. - context field_char_0 begin