# HG changeset patch # User paulson # Date 1707345582 0 # Node ID 9cde97e471df625a34fce7d4cdfed327274f5107 # Parent dafb3d343cd63c2b131bebf7d9d85e10fa55de10 Two new theorems diff -r dafb3d343cd6 -r 9cde97e471df src/HOL/Analysis/Weierstrass_Theorems.thy --- a/src/HOL/Analysis/Weierstrass_Theorems.thy Tue Feb 06 18:08:43 2024 +0000 +++ b/src/HOL/Analysis/Weierstrass_Theorems.thy Wed Feb 07 22:39:42 2024 +0000 @@ -524,7 +524,8 @@ have "\i. i \ subA \ e / real (card subA) \ 1 \ 1 - e / real (card subA) < ff i x" using e \B \ S\ ff subA(1) x by (force simp: field_split_simps) then show ?thesis - using prod_mono_strict [where f = "\x. 1 - e / card subA"] subA(2) by (force simp add: pff_def) + using prod_mono_strict[of _ subA "\x. 1 - e / card subA" ] subA + unfolding pff_def by (smt (verit, best) UN_E assms(3) subsetD) qed finally have "1 - e < pff x" . } diff -r dafb3d343cd6 -r 9cde97e471df src/HOL/Binomial.thy --- a/src/HOL/Binomial.thy Tue Feb 06 18:08:43 2024 +0000 +++ b/src/HOL/Binomial.thy Wed Feb 07 22:39:42 2024 +0000 @@ -21,6 +21,15 @@ definition binomial :: "nat \ nat \ nat" (infixl "choose" 65) where "n choose k = card {K\Pow {0.. n" shows "m choose k \ n choose k" +proof - + have "{K. K \ {0.. card K = k} \ {K. K \ {0.. card K = k}" + using assms by auto + then show ?thesis + by (simp add: binomial_def card_mono) +qed + theorem n_subsets: assumes "finite A" shows "card {B. B \ A \ card B = k} = card A choose k" diff -r dafb3d343cd6 -r 9cde97e471df src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy Tue Feb 06 18:08:43 2024 +0000 +++ b/src/HOL/Groups_Big.thy Wed Feb 07 22:39:42 2024 +0000 @@ -1606,16 +1606,23 @@ "(\i. i \ A \ 0 \ f i \ f i \ g i) \ prod f A \ prod g A" by (induct A rule: infinite_finite_induct) (force intro!: prod_nonneg mult_mono)+ +text \Only one needs to be strict\ lemma prod_mono_strict: - assumes "finite A" "\i. i \ A \ 0 \ f i \ f i < g i" "A \ {}" - shows "prod f A < prod g A" - using assms -proof (induct A rule: finite_induct) - case empty - then show ?case by simp -next - case insert - then show ?case by (force intro: mult_strict_mono' prod_nonneg) + assumes "i \ A" "f i < g i" + assumes "finite A" + assumes "\i. i \ A \ 0 \ f i \ f i \ g i" + assumes "\i. i \ A \ 0 < g i" + shows "prod f A < prod g A" +proof - + have "prod f A = f i * prod f (A - {i})" + using assms by (intro prod.remove) + also have "\ \ f i * prod g (A - {i})" + using assms by (intro mult_left_mono prod_mono) auto + also have "\ < g i * prod g (A - {i})" + using assms by (intro mult_strict_right_mono prod_pos) auto + also have "\ = prod g A" + using assms by (intro prod.remove [symmetric]) + finally show ?thesis . qed lemma prod_le_power: @@ -1742,4 +1749,24 @@ finally show ?case . qed +lemma (in linordered_semidom) prod_mono_strict': + assumes "i \ A" + assumes "finite A" + assumes "\i. i \ A \ 0 \ f i \ f i \ g i" + assumes "\i. i \ A \ 0 < g i" + assumes "f i < g i" + shows "prod f A < prod g A" +proof - + have "prod f A = f i * prod f (A - {i})" + using assms by (intro prod.remove) + also have "\ \ f i * prod g (A - {i})" + using assms by (intro mult_left_mono prod_mono) auto + also have "\ < g i * prod g (A - {i})" + using assms by (intro mult_strict_right_mono prod_pos) auto + also have "\ = prod g A" + using assms by (intro prod.remove [symmetric]) + finally show ?thesis . +qed + + end