diff --git a/src/HOL/Computational_Algebra/Factorial_Ring.thy b/src/HOL/Computational_Algebra/Factorial_Ring.thy --- a/src/HOL/Computational_Algebra/Factorial_Ring.thy +++ b/src/HOL/Computational_Algebra/Factorial_Ring.thy @@ -1,2260 +1,2260 @@ (* Title: HOL/Computational_Algebra/Factorial_Ring.thy Author: Manuel Eberl, TU Muenchen Author: Florian Haftmann, TU Muenchen *) section \Factorial (semi)rings\ theory Factorial_Ring imports Main "HOL-Library.Multiset" begin subsection \Irreducible and prime elements\ context comm_semiring_1 begin definition irreducible :: "'a \ bool" where "irreducible p \ p \ 0 \ \p dvd 1 \ (\a b. p = a * b \ a dvd 1 \ b dvd 1)" lemma not_irreducible_zero [simp]: "\irreducible 0" by (simp add: irreducible_def) lemma irreducible_not_unit: "irreducible p \ \p dvd 1" by (simp add: irreducible_def) lemma not_irreducible_one [simp]: "\irreducible 1" by (simp add: irreducible_def) lemma irreducibleI: "p \ 0 \ \p dvd 1 \ (\a b. p = a * b \ a dvd 1 \ b dvd 1) \ irreducible p" by (simp add: irreducible_def) lemma irreducibleD: "irreducible p \ p = a * b \ a dvd 1 \ b dvd 1" by (simp add: irreducible_def) lemma irreducible_mono: assumes irr: "irreducible b" and "a dvd b" "\a dvd 1" shows "irreducible a" proof (rule irreducibleI) fix c d assume "a = c * d" from assms obtain k where [simp]: "b = a * k" by auto from \a = c * d\ have "b = c * d * k" by simp hence "c dvd 1 \ (d * k) dvd 1" using irreducibleD[OF irr, of c "d * k"] by (auto simp: mult.assoc) thus "c dvd 1 \ d dvd 1" by auto qed (use assms in \auto simp: irreducible_def\) definition prime_elem :: "'a \ bool" where "prime_elem p \ p \ 0 \ \p dvd 1 \ (\a b. p dvd (a * b) \ p dvd a \ p dvd b)" lemma not_prime_elem_zero [simp]: "\prime_elem 0" by (simp add: prime_elem_def) lemma prime_elem_not_unit: "prime_elem p \ \p dvd 1" by (simp add: prime_elem_def) lemma prime_elemI: "p \ 0 \ \p dvd 1 \ (\a b. p dvd (a * b) \ p dvd a \ p dvd b) \ prime_elem p" by (simp add: prime_elem_def) lemma prime_elem_dvd_multD: "prime_elem p \ p dvd (a * b) \ p dvd a \ p dvd b" by (simp add: prime_elem_def) lemma prime_elem_dvd_mult_iff: "prime_elem p \ p dvd (a * b) \ p dvd a \ p dvd b" by (auto simp: prime_elem_def) lemma not_prime_elem_one [simp]: "\ prime_elem 1" by (auto dest: prime_elem_not_unit) lemma prime_elem_not_zeroI: assumes "prime_elem p" shows "p \ 0" using assms by (auto intro: ccontr) lemma prime_elem_dvd_power: "prime_elem p \ p dvd x ^ n \ p dvd x" by (induction n) (auto dest: prime_elem_dvd_multD intro: dvd_trans[of _ 1]) lemma prime_elem_dvd_power_iff: "prime_elem p \ n > 0 \ p dvd x ^ n \ p dvd x" by (auto dest: prime_elem_dvd_power intro: dvd_trans) lemma prime_elem_imp_nonzero [simp]: "ASSUMPTION (prime_elem x) \ x \ 0" unfolding ASSUMPTION_def by (rule prime_elem_not_zeroI) lemma prime_elem_imp_not_one [simp]: "ASSUMPTION (prime_elem x) \ x \ 1" unfolding ASSUMPTION_def by auto end lemma (in normalization_semidom) irreducible_cong: assumes "normalize a = normalize b" shows "irreducible a \ irreducible b" proof (cases "a = 0 \ a dvd 1") case True hence "\irreducible a" by (auto simp: irreducible_def) from True have "normalize a = 0 \ normalize a dvd 1" by auto also note assms finally have "b = 0 \ b dvd 1" by simp hence "\irreducible b" by (auto simp: irreducible_def) with \\irreducible a\ show ?thesis by simp next case False hence b: "b \ 0" "\is_unit b" using assms by (auto simp: is_unit_normalize[of b]) show ?thesis proof assume "irreducible a" thus "irreducible b" by (rule irreducible_mono) (use assms False b in \auto dest: associatedD2\) next assume "irreducible b" thus "irreducible a" by (rule irreducible_mono) (use assms False b in \auto dest: associatedD1\) qed qed lemma (in normalization_semidom) associatedE1: assumes "normalize a = normalize b" obtains u where "is_unit u" "a = u * b" proof (cases "a = 0") case [simp]: False from assms have [simp]: "b \ 0" by auto show ?thesis proof (rule that) show "is_unit (unit_factor a div unit_factor b)" by auto have "unit_factor a div unit_factor b * b = unit_factor a * (b div unit_factor b)" using \b \ 0\ unit_div_commute unit_div_mult_swap unit_factor_is_unit by metis also have "b div unit_factor b = normalize b" by simp finally show "a = unit_factor a div unit_factor b * b" by (metis assms unit_factor_mult_normalize) qed next case [simp]: True hence [simp]: "b = 0" using assms[symmetric] by auto show ?thesis by (intro that[of 1]) auto qed lemma (in normalization_semidom) associatedE2: assumes "normalize a = normalize b" obtains u where "is_unit u" "b = u * a" proof - from assms have "normalize b = normalize a" by simp then obtain u where "is_unit u" "b = u * a" by (elim associatedE1) thus ?thesis using that by blast qed (* TODO Move *) lemma (in normalization_semidom) normalize_power_normalize: "normalize (normalize x ^ n) = normalize (x ^ n)" proof (induction n) case (Suc n) have "normalize (normalize x ^ Suc n) = normalize (x * normalize (normalize x ^ n))" by simp also note Suc.IH finally show ?case by simp qed auto context algebraic_semidom begin lemma prime_elem_imp_irreducible: assumes "prime_elem p" shows "irreducible p" proof (rule irreducibleI) fix a b assume p_eq: "p = a * b" with assms have nz: "a \ 0" "b \ 0" by auto from p_eq have "p dvd a * b" by simp with \prime_elem p\ have "p dvd a \ p dvd b" by (rule prime_elem_dvd_multD) with \p = a * b\ have "a * b dvd 1 * b \ a * b dvd a * 1" by auto thus "a dvd 1 \ b dvd 1" by (simp only: dvd_times_left_cancel_iff[OF nz(1)] dvd_times_right_cancel_iff[OF nz(2)]) qed (insert assms, simp_all add: prime_elem_def) lemma (in algebraic_semidom) unit_imp_no_irreducible_divisors: assumes "is_unit x" "irreducible p" shows "\p dvd x" proof (rule notI) assume "p dvd x" with \is_unit x\ have "is_unit p" by (auto intro: dvd_trans) with \irreducible p\ show False by (simp add: irreducible_not_unit) qed lemma unit_imp_no_prime_divisors: assumes "is_unit x" "prime_elem p" shows "\p dvd x" using unit_imp_no_irreducible_divisors[OF assms(1) prime_elem_imp_irreducible[OF assms(2)]] . lemma prime_elem_mono: assumes "prime_elem p" "\q dvd 1" "q dvd p" shows "prime_elem q" proof - from \q dvd p\ obtain r where r: "p = q * r" by (elim dvdE) hence "p dvd q * r" by simp with \prime_elem p\ have "p dvd q \ p dvd r" by (rule prime_elem_dvd_multD) hence "p dvd q" proof assume "p dvd r" then obtain s where s: "r = p * s" by (elim dvdE) from r have "p * 1 = p * (q * s)" by (subst (asm) s) (simp add: mult_ac) with \prime_elem p\ have "q dvd 1" by (subst (asm) mult_cancel_left) auto with \\q dvd 1\ show ?thesis by contradiction qed show ?thesis proof (rule prime_elemI) fix a b assume "q dvd (a * b)" with \p dvd q\ have "p dvd (a * b)" by (rule dvd_trans) with \prime_elem p\ have "p dvd a \ p dvd b" by (rule prime_elem_dvd_multD) with \q dvd p\ show "q dvd a \ q dvd b" by (blast intro: dvd_trans) qed (insert assms, auto) qed lemma irreducibleD': assumes "irreducible a" "b dvd a" shows "a dvd b \ is_unit b" proof - from assms obtain c where c: "a = b * c" by (elim dvdE) from irreducibleD[OF assms(1) this] have "is_unit b \ is_unit c" . thus ?thesis by (auto simp: c mult_unit_dvd_iff) qed lemma irreducibleI': assumes "a \ 0" "\is_unit a" "\b. b dvd a \ a dvd b \ is_unit b" shows "irreducible a" proof (rule irreducibleI) fix b c assume a_eq: "a = b * c" hence "a dvd b \ is_unit b" by (intro assms) simp_all thus "is_unit b \ is_unit c" proof assume "a dvd b" hence "b * c dvd b * 1" by (simp add: a_eq) moreover from \a \ 0\ a_eq have "b \ 0" by auto ultimately show ?thesis by (subst (asm) dvd_times_left_cancel_iff) auto qed blast qed (simp_all add: assms(1,2)) lemma irreducible_altdef: "irreducible x \ x \ 0 \ \is_unit x \ (\b. b dvd x \ x dvd b \ is_unit b)" using irreducibleI'[of x] irreducibleD'[of x] irreducible_not_unit[of x] by auto lemma prime_elem_multD: assumes "prime_elem (a * b)" shows "is_unit a \ is_unit b" proof - from assms have "a \ 0" "b \ 0" by (auto dest!: prime_elem_not_zeroI) moreover from assms prime_elem_dvd_multD [of "a * b"] have "a * b dvd a \ a * b dvd b" by auto ultimately show ?thesis using dvd_times_left_cancel_iff [of a b 1] dvd_times_right_cancel_iff [of b a 1] by auto qed lemma prime_elemD2: assumes "prime_elem p" and "a dvd p" and "\ is_unit a" shows "p dvd a" proof - from \a dvd p\ obtain b where "p = a * b" .. with \prime_elem p\ prime_elem_multD \\ is_unit a\ have "is_unit b" by auto with \p = a * b\ show ?thesis by (auto simp add: mult_unit_dvd_iff) qed lemma prime_elem_dvd_prod_msetE: assumes "prime_elem p" assumes dvd: "p dvd prod_mset A" obtains a where "a \# A" and "p dvd a" proof - from dvd have "\a. a \# A \ p dvd a" proof (induct A) case empty then show ?case using \prime_elem p\ by (simp add: prime_elem_not_unit) next case (add a A) then have "p dvd a * prod_mset A" by simp with \prime_elem p\ consider (A) "p dvd prod_mset A" | (B) "p dvd a" by (blast dest: prime_elem_dvd_multD) then show ?case proof cases case B then show ?thesis by auto next case A with add.hyps obtain b where "b \# A" "p dvd b" by auto then show ?thesis by auto qed qed with that show thesis by blast qed context begin lemma prime_elem_powerD: assumes "prime_elem (p ^ n)" shows "prime_elem p \ n = 1" proof (cases n) case (Suc m) note assms also from Suc have "p ^ n = p * p^m" by simp finally have "is_unit p \ is_unit (p^m)" by (rule prime_elem_multD) moreover from assms have "\is_unit p" by (simp add: prime_elem_def is_unit_power_iff) ultimately have "is_unit (p ^ m)" by simp with \\is_unit p\ have "m = 0" by (simp add: is_unit_power_iff) with Suc assms show ?thesis by simp qed (insert assms, simp_all) lemma prime_elem_power_iff: "prime_elem (p ^ n) \ prime_elem p \ n = 1" by (auto dest: prime_elem_powerD) end lemma irreducible_mult_unit_left: "is_unit a \ irreducible (a * p) \ irreducible p" by (auto simp: irreducible_altdef mult.commute[of a] is_unit_mult_iff mult_unit_dvd_iff dvd_mult_unit_iff) lemma prime_elem_mult_unit_left: "is_unit a \ prime_elem (a * p) \ prime_elem p" by (auto simp: prime_elem_def mult.commute[of a] is_unit_mult_iff mult_unit_dvd_iff) lemma prime_elem_dvd_cases: assumes pk: "p*k dvd m*n" and p: "prime_elem p" shows "(\x. k dvd x*n \ m = p*x) \ (\y. k dvd m*y \ n = p*y)" proof - have "p dvd m*n" using dvd_mult_left pk by blast then consider "p dvd m" | "p dvd n" using p prime_elem_dvd_mult_iff by blast then show ?thesis proof cases case 1 then obtain a where "m = p * a" by (metis dvd_mult_div_cancel) then have "\x. k dvd x * n \ m = p * x" using p pk by (auto simp: mult.assoc) then show ?thesis .. next case 2 then obtain b where "n = p * b" by (metis dvd_mult_div_cancel) with p pk have "\y. k dvd m*y \ n = p*y" by (metis dvd_mult_right dvd_times_left_cancel_iff mult.left_commute mult_zero_left) then show ?thesis .. qed qed lemma prime_elem_power_dvd_prod: assumes pc: "p^c dvd m*n" and p: "prime_elem p" shows "\a b. a+b = c \ p^a dvd m \ p^b dvd n" using pc proof (induct c arbitrary: m n) case 0 show ?case by simp next case (Suc c) consider x where "p^c dvd x*n" "m = p*x" | y where "p^c dvd m*y" "n = p*y" using prime_elem_dvd_cases [of _ "p^c", OF _ p] Suc.prems by force then show ?case proof cases case (1 x) with Suc.hyps[of x n] obtain a b where "a + b = c \ p ^ a dvd x \ p ^ b dvd n" by blast with 1 have "Suc a + b = Suc c \ p ^ Suc a dvd m \ p ^ b dvd n" by (auto intro: mult_dvd_mono) thus ?thesis by blast next case (2 y) with Suc.hyps[of m y] obtain a b where "a + b = c \ p ^ a dvd m \ p ^ b dvd y" by blast with 2 have "a + Suc b = Suc c \ p ^ a dvd m \ p ^ Suc b dvd n" by (auto intro: mult_dvd_mono) with Suc.hyps [of m y] show "\a b. a + b = Suc c \ p ^ a dvd m \ p ^ b dvd n" by blast qed qed lemma prime_elem_power_dvd_cases: assumes "p ^ c dvd m * n" and "a + b = Suc c" and "prime_elem p" shows "p ^ a dvd m \ p ^ b dvd n" proof - from assms obtain r s where "r + s = c \ p ^ r dvd m \ p ^ s dvd n" by (blast dest: prime_elem_power_dvd_prod) moreover with assms have "a \ r \ b \ s" by arith ultimately show ?thesis by (auto intro: power_le_dvd) qed lemma prime_elem_not_unit' [simp]: "ASSUMPTION (prime_elem x) \ \is_unit x" unfolding ASSUMPTION_def by (rule prime_elem_not_unit) lemma prime_elem_dvd_power_iff: assumes "prime_elem p" shows "p dvd a ^ n \ p dvd a \ n > 0" using assms by (induct n) (auto dest: prime_elem_not_unit prime_elem_dvd_multD) lemma prime_power_dvd_multD: assumes "prime_elem p" assumes "p ^ n dvd a * b" and "n > 0" and "\ p dvd a" shows "p ^ n dvd b" using \p ^ n dvd a * b\ and \n > 0\ proof (induct n arbitrary: b) case 0 then show ?case by simp next case (Suc n) show ?case proof (cases "n = 0") case True with Suc \prime_elem p\ \\ p dvd a\ show ?thesis by (simp add: prime_elem_dvd_mult_iff) next case False then have "n > 0" by simp from \prime_elem p\ have "p \ 0" by auto from Suc.prems have *: "p * p ^ n dvd a * b" by simp then have "p dvd a * b" by (rule dvd_mult_left) with Suc \prime_elem p\ \\ p dvd a\ have "p dvd b" by (simp add: prime_elem_dvd_mult_iff) moreover define c where "c = b div p" ultimately have b: "b = p * c" by simp with * have "p * p ^ n dvd p * (a * c)" by (simp add: ac_simps) with \p \ 0\ have "p ^ n dvd a * c" by simp with Suc.hyps \n > 0\ have "p ^ n dvd c" by blast with \p \ 0\ show ?thesis by (simp add: b) qed qed end subsection \Generalized primes: normalized prime elements\ context normalization_semidom begin lemma irreducible_normalized_divisors: assumes "irreducible x" "y dvd x" "normalize y = y" shows "y = 1 \ y = normalize x" proof - from assms have "is_unit y \ x dvd y" by (auto simp: irreducible_altdef) thus ?thesis proof (elim disjE) assume "is_unit y" hence "normalize y = 1" by (simp add: is_unit_normalize) with assms show ?thesis by simp next assume "x dvd y" with \y dvd x\ have "normalize y = normalize x" by (rule associatedI) with assms show ?thesis by simp qed qed lemma irreducible_normalize_iff [simp]: "irreducible (normalize x) = irreducible x" using irreducible_mult_unit_left[of "1 div unit_factor x" x] by (cases "x = 0") (simp_all add: unit_div_commute) lemma prime_elem_normalize_iff [simp]: "prime_elem (normalize x) = prime_elem x" using prime_elem_mult_unit_left[of "1 div unit_factor x" x] by (cases "x = 0") (simp_all add: unit_div_commute) lemma prime_elem_associated: assumes "prime_elem p" and "prime_elem q" and "q dvd p" shows "normalize q = normalize p" using \q dvd p\ proof (rule associatedI) from \prime_elem q\ have "\ is_unit q" by (auto simp add: prime_elem_not_unit) with \prime_elem p\ \q dvd p\ show "p dvd q" by (blast intro: prime_elemD2) qed definition prime :: "'a \ bool" where "prime p \ prime_elem p \ normalize p = p" lemma not_prime_0 [simp]: "\prime 0" by (simp add: prime_def) lemma not_prime_unit: "is_unit x \ \prime x" using prime_elem_not_unit[of x] by (auto simp add: prime_def) lemma not_prime_1 [simp]: "\prime 1" by (simp add: not_prime_unit) lemma primeI: "prime_elem x \ normalize x = x \ prime x" by (simp add: prime_def) lemma prime_imp_prime_elem [dest]: "prime p \ prime_elem p" by (simp add: prime_def) lemma normalize_prime: "prime p \ normalize p = p" by (simp add: prime_def) lemma prime_normalize_iff [simp]: "prime (normalize p) \ prime_elem p" by (auto simp add: prime_def) lemma prime_power_iff: "prime (p ^ n) \ prime p \ n = 1" by (auto simp: prime_def prime_elem_power_iff) lemma prime_imp_nonzero [simp]: "ASSUMPTION (prime x) \ x \ 0" unfolding ASSUMPTION_def prime_def by auto lemma prime_imp_not_one [simp]: "ASSUMPTION (prime x) \ x \ 1" unfolding ASSUMPTION_def by auto lemma prime_not_unit' [simp]: "ASSUMPTION (prime x) \ \is_unit x" unfolding ASSUMPTION_def prime_def by auto lemma prime_normalize' [simp]: "ASSUMPTION (prime x) \ normalize x = x" unfolding ASSUMPTION_def prime_def by simp lemma unit_factor_prime: "prime x \ unit_factor x = 1" using unit_factor_normalize[of x] unfolding prime_def by auto lemma unit_factor_prime' [simp]: "ASSUMPTION (prime x) \ unit_factor x = 1" unfolding ASSUMPTION_def by (rule unit_factor_prime) lemma prime_imp_prime_elem' [simp]: "ASSUMPTION (prime x) \ prime_elem x" by (simp add: prime_def ASSUMPTION_def) lemma prime_dvd_multD: "prime p \ p dvd a * b \ p dvd a \ p dvd b" by (intro prime_elem_dvd_multD) simp_all lemma prime_dvd_mult_iff: "prime p \ p dvd a * b \ p dvd a \ p dvd b" by (auto dest: prime_dvd_multD) lemma prime_dvd_power: "prime p \ p dvd x ^ n \ p dvd x" by (auto dest!: prime_elem_dvd_power simp: prime_def) lemma prime_dvd_power_iff: "prime p \ n > 0 \ p dvd x ^ n \ p dvd x" by (subst prime_elem_dvd_power_iff) simp_all lemma prime_dvd_prod_mset_iff: "prime p \ p dvd prod_mset A \ (\x. x \# A \ p dvd x)" by (induction A) (simp_all add: prime_elem_dvd_mult_iff prime_imp_prime_elem, blast+) lemma prime_dvd_prod_iff: "finite A \ prime p \ p dvd prod f A \ (\x\A. p dvd f x)" by (auto simp: prime_dvd_prod_mset_iff prod_unfold_prod_mset) lemma primes_dvd_imp_eq: assumes "prime p" "prime q" "p dvd q" shows "p = q" proof - from assms have "irreducible q" by (simp add: prime_elem_imp_irreducible prime_def) from irreducibleD'[OF this \p dvd q\] assms have "q dvd p" by simp with \p dvd q\ have "normalize p = normalize q" by (rule associatedI) with assms show "p = q" by simp qed lemma prime_dvd_prod_mset_primes_iff: assumes "prime p" "\q. q \# A \ prime q" shows "p dvd prod_mset A \ p \# A" proof - from assms(1) have "p dvd prod_mset A \ (\x. x \# A \ p dvd x)" by (rule prime_dvd_prod_mset_iff) also from assms have "\ \ p \# A" by (auto dest: primes_dvd_imp_eq) finally show ?thesis . qed lemma prod_mset_primes_dvd_imp_subset: assumes "prod_mset A dvd prod_mset B" "\p. p \# A \ prime p" "\p. p \# B \ prime p" shows "A \# B" using assms proof (induction A arbitrary: B) case empty thus ?case by simp next case (add p A B) hence p: "prime p" by simp define B' where "B' = B - {#p#}" from add.prems have "p dvd prod_mset B" by (simp add: dvd_mult_left) with add.prems have "p \# B" by (subst (asm) (2) prime_dvd_prod_mset_primes_iff) simp_all hence B: "B = B' + {#p#}" by (simp add: B'_def) from add.prems p have "A \# B'" by (intro add.IH) (simp_all add: B) thus ?case by (simp add: B) qed lemma prod_mset_dvd_prod_mset_primes_iff: assumes "\x. x \# A \ prime x" "\x. x \# B \ prime x" shows "prod_mset A dvd prod_mset B \ A \# B" using assms by (auto intro: prod_mset_subset_imp_dvd prod_mset_primes_dvd_imp_subset) lemma is_unit_prod_mset_primes_iff: assumes "\x. x \# A \ prime x" shows "is_unit (prod_mset A) \ A = {#}" by (auto simp add: is_unit_prod_mset_iff) (meson all_not_in_conv assms not_prime_unit set_mset_eq_empty_iff) lemma prod_mset_primes_irreducible_imp_prime: assumes irred: "irreducible (prod_mset A)" assumes A: "\x. x \# A \ prime x" assumes B: "\x. x \# B \ prime x" assumes C: "\x. x \# C \ prime x" assumes dvd: "prod_mset A dvd prod_mset B * prod_mset C" shows "prod_mset A dvd prod_mset B \ prod_mset A dvd prod_mset C" proof - from dvd have "prod_mset A dvd prod_mset (B + C)" by simp with A B C have subset: "A \# B + C" by (subst (asm) prod_mset_dvd_prod_mset_primes_iff) auto define A1 and A2 where "A1 = A \# B" and "A2 = A - A1" have "A = A1 + A2" unfolding A1_def A2_def by (rule sym, intro subset_mset.add_diff_inverse) simp_all from subset have "A1 \# B" "A2 \# C" by (auto simp: A1_def A2_def Multiset.subset_eq_diff_conv Multiset.union_commute) from \A = A1 + A2\ have "prod_mset A = prod_mset A1 * prod_mset A2" by simp from irred and this have "is_unit (prod_mset A1) \ is_unit (prod_mset A2)" by (rule irreducibleD) with A have "A1 = {#} \ A2 = {#}" unfolding A1_def A2_def by (subst (asm) (1 2) is_unit_prod_mset_primes_iff) (auto dest: Multiset.in_diffD) with dvd \A = A1 + A2\ \A1 \# B\ \A2 \# C\ show ?thesis by (auto intro: prod_mset_subset_imp_dvd) qed lemma prod_mset_primes_finite_divisor_powers: assumes A: "\x. x \# A \ prime x" assumes B: "\x. x \# B \ prime x" assumes "A \ {#}" shows "finite {n. prod_mset A ^ n dvd prod_mset B}" proof - from \A \ {#}\ obtain x where x: "x \# A" by blast define m where "m = count B x" have "{n. prod_mset A ^ n dvd prod_mset B} \ {..m}" proof safe fix n assume dvd: "prod_mset A ^ n dvd prod_mset B" from x have "x ^ n dvd prod_mset A ^ n" by (intro dvd_power_same dvd_prod_mset) also note dvd also have "x ^ n = prod_mset (replicate_mset n x)" by simp finally have "replicate_mset n x \# B" by (rule prod_mset_primes_dvd_imp_subset) (insert A B x, simp_all split: if_splits) thus "n \ m" by (simp add: count_le_replicate_mset_subset_eq m_def) qed moreover have "finite {..m}" by simp ultimately show ?thesis by (rule finite_subset) qed end subsection \In a semiring with GCD, each irreducible element is a prime element\ context semiring_gcd begin lemma irreducible_imp_prime_elem_gcd: assumes "irreducible x" shows "prime_elem x" proof (rule prime_elemI) fix a b assume "x dvd a * b" from dvd_productE[OF this] obtain y z where yz: "x = y * z" "y dvd a" "z dvd b" . from \irreducible x\ and \x = y * z\ have "is_unit y \ is_unit z" by (rule irreducibleD) with yz show "x dvd a \ x dvd b" by (auto simp: mult_unit_dvd_iff mult_unit_dvd_iff') qed (insert assms, auto simp: irreducible_not_unit) lemma prime_elem_imp_coprime: assumes "prime_elem p" "\p dvd n" shows "coprime p n" proof (rule coprimeI) fix d assume "d dvd p" "d dvd n" show "is_unit d" proof (rule ccontr) assume "\is_unit d" from \prime_elem p\ and \d dvd p\ and this have "p dvd d" by (rule prime_elemD2) from this and \d dvd n\ have "p dvd n" by (rule dvd_trans) with \\p dvd n\ show False by contradiction qed qed lemma prime_imp_coprime: assumes "prime p" "\p dvd n" shows "coprime p n" using assms by (simp add: prime_elem_imp_coprime) lemma prime_elem_imp_power_coprime: "prime_elem p \ \ p dvd a \ coprime a (p ^ m)" by (cases "m > 0") (auto dest: prime_elem_imp_coprime simp add: ac_simps) lemma prime_imp_power_coprime: "prime p \ \ p dvd a \ coprime a (p ^ m)" by (rule prime_elem_imp_power_coprime) simp_all lemma prime_elem_divprod_pow: assumes p: "prime_elem p" and ab: "coprime a b" and pab: "p^n dvd a * b" shows "p^n dvd a \ p^n dvd b" using assms proof - from p have "\ is_unit p" by simp with ab p have "\ p dvd a \ \ p dvd b" using not_coprimeI by blast with p have "coprime (p ^ n) a \ coprime (p ^ n) b" by (auto dest: prime_elem_imp_power_coprime simp add: ac_simps) with pab show ?thesis by (auto simp add: coprime_dvd_mult_left_iff coprime_dvd_mult_right_iff) qed lemma primes_coprime: "prime p \ prime q \ p \ q \ coprime p q" using prime_imp_coprime primes_dvd_imp_eq by blast end subsection \Factorial semirings: algebraic structures with unique prime factorizations\ class factorial_semiring = normalization_semidom + assumes prime_factorization_exists: "x \ 0 \ \A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize x" text \Alternative characterization\ lemma (in normalization_semidom) factorial_semiring_altI_aux: assumes finite_divisors: "\x. x \ 0 \ finite {y. y dvd x \ normalize y = y}" assumes irreducible_imp_prime_elem: "\x. irreducible x \ prime_elem x" assumes "x \ 0" shows "\A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize x" using \x \ 0\ proof (induction "card {b. b dvd x \ normalize b = b}" arbitrary: x rule: less_induct) case (less a) let ?fctrs = "\a. {b. b dvd a \ normalize b = b}" show ?case proof (cases "is_unit a") case True thus ?thesis by (intro exI[of _ "{#}"]) (auto simp: is_unit_normalize) next case False show ?thesis proof (cases "\b. b dvd a \ \is_unit b \ \a dvd b") case False with \\is_unit a\ less.prems have "irreducible a" by (auto simp: irreducible_altdef) hence "prime_elem a" by (rule irreducible_imp_prime_elem) thus ?thesis by (intro exI[of _ "{#normalize a#}"]) auto next case True then obtain b where b: "b dvd a" "\ is_unit b" "\ a dvd b" by auto from b have "?fctrs b \ ?fctrs a" by (auto intro: dvd_trans) moreover from b have "normalize a \ ?fctrs b" "normalize a \ ?fctrs a" by simp_all hence "?fctrs b \ ?fctrs a" by blast ultimately have "?fctrs b \ ?fctrs a" by (subst subset_not_subset_eq) blast with finite_divisors[OF \a \ 0\] have "card (?fctrs b) < card (?fctrs a)" by (rule psubset_card_mono) moreover from \a \ 0\ b have "b \ 0" by auto ultimately have "\A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize b" by (intro less) auto then obtain A where A: "(\x. x \# A \ prime_elem x) \ normalize (\\<^sub># A) = normalize b" by auto define c where "c = a div b" from b have c: "a = b * c" by (simp add: c_def) from less.prems c have "c \ 0" by auto from b c have "?fctrs c \ ?fctrs a" by (auto intro: dvd_trans) moreover have "normalize a \ ?fctrs c" proof safe assume "normalize a dvd c" hence "b * c dvd 1 * c" by (simp add: c) hence "b dvd 1" by (subst (asm) dvd_times_right_cancel_iff) fact+ with b show False by simp qed with \normalize a \ ?fctrs a\ have "?fctrs a \ ?fctrs c" by blast ultimately have "?fctrs c \ ?fctrs a" by (subst subset_not_subset_eq) blast with finite_divisors[OF \a \ 0\] have "card (?fctrs c) < card (?fctrs a)" by (rule psubset_card_mono) with \c \ 0\ have "\A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize c" by (intro less) auto then obtain B where B: "(\x. x \# B \ prime_elem x) \ normalize (\\<^sub># B) = normalize c" by auto show ?thesis proof (rule exI[of _ "A + B"]; safe) have "normalize (prod_mset (A + B)) = normalize (normalize (prod_mset A) * normalize (prod_mset B))" by simp also have "\ = normalize (b * c)" by (simp only: A B) auto also have "b * c = a" using c by simp finally show "normalize (prod_mset (A + B)) = normalize a" . next qed (use A B in auto) qed qed qed lemma factorial_semiring_altI: assumes finite_divisors: "\x::'a. x \ 0 \ finite {y. y dvd x \ normalize y = y}" assumes irreducible_imp_prime: "\x::'a. irreducible x \ prime_elem x" shows "OFCLASS('a :: normalization_semidom, factorial_semiring_class)" by intro_classes (rule factorial_semiring_altI_aux[OF assms]) text \Properties\ context factorial_semiring begin lemma prime_factorization_exists': assumes "x \ 0" obtains A where "\x. x \# A \ prime x" "normalize (prod_mset A) = normalize x" proof - from prime_factorization_exists[OF assms] obtain A where A: "\x. x \# A \ prime_elem x" "normalize (prod_mset A) = normalize x" by blast define A' where "A' = image_mset normalize A" have "normalize (prod_mset A') = normalize (prod_mset A)" by (simp add: A'_def normalize_prod_mset_normalize) also note A(2) finally have "normalize (prod_mset A') = normalize x" by simp moreover from A(1) have "\x. x \# A' \ prime x" by (auto simp: prime_def A'_def) ultimately show ?thesis by (intro that[of A']) blast qed lemma irreducible_imp_prime_elem: assumes "irreducible x" shows "prime_elem x" proof (rule prime_elemI) fix a b assume dvd: "x dvd a * b" from assms have "x \ 0" by auto show "x dvd a \ x dvd b" proof (cases "a = 0 \ b = 0") case False hence "a \ 0" "b \ 0" by blast+ note nz = \x \ 0\ this from nz[THEN prime_factorization_exists'] obtain A B C where ABC: "\z. z \# A \ prime z" "normalize (\\<^sub># A) = normalize x" "\z. z \# B \ prime z" "normalize (\\<^sub># B) = normalize a" "\z. z \# C \ prime z" "normalize (\\<^sub># C) = normalize b" by this blast have "irreducible (prod_mset A)" by (subst irreducible_cong[OF ABC(2)]) fact moreover have "normalize (prod_mset A) dvd normalize (normalize (prod_mset B) * normalize (prod_mset C))" unfolding ABC using dvd by simp hence "prod_mset A dvd prod_mset B * prod_mset C" unfolding normalize_mult_normalize_left normalize_mult_normalize_right by simp ultimately have "prod_mset A dvd prod_mset B \ prod_mset A dvd prod_mset C" by (intro prod_mset_primes_irreducible_imp_prime) (use ABC in auto) hence "normalize (prod_mset A) dvd normalize (prod_mset B) \ normalize (prod_mset A) dvd normalize (prod_mset C)" by simp thus ?thesis unfolding ABC by simp qed auto qed (use assms in \simp_all add: irreducible_def\) lemma finite_divisor_powers: assumes "y \ 0" "\is_unit x" shows "finite {n. x ^ n dvd y}" proof (cases "x = 0") case True with assms have "{n. x ^ n dvd y} = {0}" by (auto simp: power_0_left) thus ?thesis by simp next case False note nz = this \y \ 0\ from nz[THEN prime_factorization_exists'] obtain A B where AB: "\z. z \# A \ prime z" "normalize (\\<^sub># A) = normalize x" "\z. z \# B \ prime z" "normalize (\\<^sub># B) = normalize y" by this blast from AB assms have "A \ {#}" by (auto simp: normalize_1_iff) from AB(2,4) prod_mset_primes_finite_divisor_powers [of A B, OF AB(1,3) this] have "finite {n. prod_mset A ^ n dvd prod_mset B}" by simp also have "{n. prod_mset A ^ n dvd prod_mset B} = {n. normalize (normalize (prod_mset A) ^ n) dvd normalize (prod_mset B)}" unfolding normalize_power_normalize by simp also have "\ = {n. x ^ n dvd y}" unfolding AB unfolding normalize_power_normalize by simp finally show ?thesis . qed lemma finite_prime_divisors: assumes "x \ 0" shows "finite {p. prime p \ p dvd x}" proof - from prime_factorization_exists'[OF assms] obtain A where A: "\z. z \# A \ prime z" "normalize (\\<^sub># A) = normalize x" by this blast have "{p. prime p \ p dvd x} \ set_mset A" proof safe fix p assume p: "prime p" and dvd: "p dvd x" from dvd have "p dvd normalize x" by simp also from A have "normalize x = normalize (prod_mset A)" by simp finally have "p dvd prod_mset A" by simp thus "p \# A" using p A by (subst (asm) prime_dvd_prod_mset_primes_iff) qed moreover have "finite (set_mset A)" by simp ultimately show ?thesis by (rule finite_subset) qed lemma infinite_unit_divisor_powers: assumes "y \ 0" assumes "is_unit x" shows "infinite {n. x^n dvd y}" proof - - from `is_unit x` have "is_unit (x^n)" for n + from \is_unit x\ have "is_unit (x^n)" for n using is_unit_power_iff by auto hence "x^n dvd y" for n by auto hence "{n. x^n dvd y} = UNIV" by auto thus ?thesis by auto qed corollary is_unit_iff_infinite_divisor_powers: assumes "y \ 0" shows "is_unit x \ infinite {n. x^n dvd y}" using infinite_unit_divisor_powers finite_divisor_powers assms by auto lemma prime_elem_iff_irreducible: "prime_elem x \ irreducible x" by (blast intro: irreducible_imp_prime_elem prime_elem_imp_irreducible) lemma prime_divisor_exists: assumes "a \ 0" "\is_unit a" shows "\b. b dvd a \ prime b" proof - from prime_factorization_exists'[OF assms(1)] obtain A where A: "\z. z \# A \ prime z" "normalize (\\<^sub># A) = normalize a" by this blast with assms have "A \ {#}" by auto then obtain x where "x \# A" by blast with A(1) have *: "x dvd normalize (prod_mset A)" "prime x" by (auto simp: dvd_prod_mset) hence "x dvd a" by (simp add: A(2)) with * show ?thesis by blast qed lemma prime_divisors_induct [case_names zero unit factor]: assumes "P 0" "\x. is_unit x \ P x" "\p x. prime p \ P x \ P (p * x)" shows "P x" proof (cases "x = 0") case False from prime_factorization_exists'[OF this] obtain A where A: "\z. z \# A \ prime z" "normalize (\\<^sub># A) = normalize x" by this blast from A obtain u where u: "is_unit u" "x = u * prod_mset A" by (elim associatedE2) from A(1) have "P (u * prod_mset A)" proof (induction A) case (add p A) from add.prems have "prime p" by simp moreover from add.prems have "P (u * prod_mset A)" by (intro add.IH) simp_all ultimately have "P (p * (u * prod_mset A))" by (rule assms(3)) thus ?case by (simp add: mult_ac) qed (simp_all add: assms False u) with A u show ?thesis by simp qed (simp_all add: assms(1)) lemma no_prime_divisors_imp_unit: assumes "a \ 0" "\b. b dvd a \ normalize b = b \ \ prime_elem b" shows "is_unit a" proof (rule ccontr) assume "\is_unit a" from prime_divisor_exists[OF assms(1) this] obtain b where "b dvd a" "prime b" by auto with assms(2)[of b] show False by (simp add: prime_def) qed lemma prime_divisorE: assumes "a \ 0" and "\ is_unit a" obtains p where "prime p" and "p dvd a" using assms no_prime_divisors_imp_unit unfolding prime_def by blast definition multiplicity :: "'a \ 'a \ nat" where "multiplicity p x = (if finite {n. p ^ n dvd x} then Max {n. p ^ n dvd x} else 0)" lemma multiplicity_dvd: "p ^ multiplicity p x dvd x" proof (cases "finite {n. p ^ n dvd x}") case True hence "multiplicity p x = Max {n. p ^ n dvd x}" by (simp add: multiplicity_def) also have "\ \ {n. p ^ n dvd x}" by (rule Max_in) (auto intro!: True exI[of _ "0::nat"]) finally show ?thesis by simp qed (simp add: multiplicity_def) lemma multiplicity_dvd': "n \ multiplicity p x \ p ^ n dvd x" by (rule dvd_trans[OF le_imp_power_dvd multiplicity_dvd]) context fixes x p :: 'a assumes xp: "x \ 0" "\is_unit p" begin lemma multiplicity_eq_Max: "multiplicity p x = Max {n. p ^ n dvd x}" using finite_divisor_powers[OF xp] by (simp add: multiplicity_def) lemma multiplicity_geI: assumes "p ^ n dvd x" shows "multiplicity p x \ n" proof - from assms have "n \ Max {n. p ^ n dvd x}" by (intro Max_ge finite_divisor_powers xp) simp_all thus ?thesis by (subst multiplicity_eq_Max) qed lemma multiplicity_lessI: assumes "\p ^ n dvd x" shows "multiplicity p x < n" proof (rule ccontr) assume "\(n > multiplicity p x)" hence "p ^ n dvd x" by (intro multiplicity_dvd') simp with assms show False by contradiction qed lemma power_dvd_iff_le_multiplicity: "p ^ n dvd x \ n \ multiplicity p x" using multiplicity_geI[of n] multiplicity_lessI[of n] by (cases "p ^ n dvd x") auto lemma multiplicity_eq_zero_iff: shows "multiplicity p x = 0 \ \p dvd x" using power_dvd_iff_le_multiplicity[of 1] by auto lemma multiplicity_gt_zero_iff: shows "multiplicity p x > 0 \ p dvd x" using power_dvd_iff_le_multiplicity[of 1] by auto lemma multiplicity_decompose: "\p dvd (x div p ^ multiplicity p x)" proof assume *: "p dvd x div p ^ multiplicity p x" have "x = x div p ^ multiplicity p x * (p ^ multiplicity p x)" using multiplicity_dvd[of p x] by simp also from * have "x div p ^ multiplicity p x = (x div p ^ multiplicity p x div p) * p" by simp also have "x div p ^ multiplicity p x div p * p * p ^ multiplicity p x = x div p ^ multiplicity p x div p * p ^ Suc (multiplicity p x)" by (simp add: mult_assoc) also have "p ^ Suc (multiplicity p x) dvd \" by (rule dvd_triv_right) finally show False by (subst (asm) power_dvd_iff_le_multiplicity) simp qed lemma multiplicity_decompose': obtains y where "x = p ^ multiplicity p x * y" "\p dvd y" using that[of "x div p ^ multiplicity p x"] by (simp add: multiplicity_decompose multiplicity_dvd) end lemma multiplicity_zero [simp]: "multiplicity p 0 = 0" by (simp add: multiplicity_def) lemma prime_elem_multiplicity_eq_zero_iff: "prime_elem p \ x \ 0 \ multiplicity p x = 0 \ \p dvd x" by (rule multiplicity_eq_zero_iff) simp_all lemma prime_multiplicity_other: assumes "prime p" "prime q" "p \ q" shows "multiplicity p q = 0" using assms by (subst prime_elem_multiplicity_eq_zero_iff) (auto dest: primes_dvd_imp_eq) lemma prime_multiplicity_gt_zero_iff: "prime_elem p \ x \ 0 \ multiplicity p x > 0 \ p dvd x" by (rule multiplicity_gt_zero_iff) simp_all lemma multiplicity_unit_left: "is_unit p \ multiplicity p x = 0" by (simp add: multiplicity_def is_unit_power_iff unit_imp_dvd) lemma multiplicity_unit_right: assumes "is_unit x" shows "multiplicity p x = 0" proof (cases "is_unit p \ x = 0") case False with multiplicity_lessI[of x p 1] this assms show ?thesis by (auto dest: dvd_unit_imp_unit) qed (auto simp: multiplicity_unit_left) lemma multiplicity_one [simp]: "multiplicity p 1 = 0" by (rule multiplicity_unit_right) simp_all lemma multiplicity_eqI: assumes "p ^ n dvd x" "\p ^ Suc n dvd x" shows "multiplicity p x = n" proof - consider "x = 0" | "is_unit p" | "x \ 0" "\is_unit p" by blast thus ?thesis proof cases assume xp: "x \ 0" "\is_unit p" from xp assms(1) have "multiplicity p x \ n" by (intro multiplicity_geI) moreover from assms(2) xp have "multiplicity p x < Suc n" by (intro multiplicity_lessI) ultimately show ?thesis by simp next assume "is_unit p" hence "is_unit (p ^ Suc n)" by (simp add: is_unit_power_iff del: power_Suc) hence "p ^ Suc n dvd x" by (rule unit_imp_dvd) with \\p ^ Suc n dvd x\ show ?thesis by contradiction qed (insert assms, simp_all) qed context fixes x p :: 'a assumes xp: "x \ 0" "\is_unit p" begin lemma multiplicity_times_same: assumes "p \ 0" shows "multiplicity p (p * x) = Suc (multiplicity p x)" proof (rule multiplicity_eqI) show "p ^ Suc (multiplicity p x) dvd p * x" by (auto intro!: mult_dvd_mono multiplicity_dvd) from xp assms show "\ p ^ Suc (Suc (multiplicity p x)) dvd p * x" using power_dvd_iff_le_multiplicity[OF xp, of "Suc (multiplicity p x)"] by simp qed end lemma multiplicity_same_power': "multiplicity p (p ^ n) = (if p = 0 \ is_unit p then 0 else n)" proof - consider "p = 0" | "is_unit p" |"p \ 0" "\is_unit p" by blast thus ?thesis proof cases assume "p \ 0" "\is_unit p" thus ?thesis by (induction n) (simp_all add: multiplicity_times_same) qed (simp_all add: power_0_left multiplicity_unit_left) qed lemma multiplicity_same_power: "p \ 0 \ \is_unit p \ multiplicity p (p ^ n) = n" by (simp add: multiplicity_same_power') lemma multiplicity_prime_elem_times_other: assumes "prime_elem p" "\p dvd q" shows "multiplicity p (q * x) = multiplicity p x" proof (cases "x = 0") case False show ?thesis proof (rule multiplicity_eqI) have "1 * p ^ multiplicity p x dvd q * x" by (intro mult_dvd_mono multiplicity_dvd) simp_all thus "p ^ multiplicity p x dvd q * x" by simp next define n where "n = multiplicity p x" from assms have "\is_unit p" by simp from multiplicity_decompose'[OF False this] obtain y where y [folded n_def]: "x = p ^ multiplicity p x * y" "\ p dvd y" . from y have "p ^ Suc n dvd q * x \ p ^ n * p dvd p ^ n * (q * y)" by (simp add: mult_ac) also from assms have "\ \ p dvd q * y" by simp also have "\ \ p dvd q \ p dvd y" by (rule prime_elem_dvd_mult_iff) fact+ also from assms y have "\ \ False" by simp finally show "\(p ^ Suc n dvd q * x)" by blast qed qed simp_all lemma multiplicity_self: assumes "p \ 0" "\is_unit p" shows "multiplicity p p = 1" proof - from assms have "multiplicity p p = Max {n. p ^ n dvd p}" by (simp add: multiplicity_eq_Max) also from assms have "p ^ n dvd p \ n \ 1" for n using dvd_power_iff[of p n 1] by auto hence "{n. p ^ n dvd p} = {..1}" by auto also have "\ = {0,1}" by auto finally show ?thesis by simp qed lemma multiplicity_times_unit_left: assumes "is_unit c" shows "multiplicity (c * p) x = multiplicity p x" proof - from assms have "{n. (c * p) ^ n dvd x} = {n. p ^ n dvd x}" by (subst mult.commute) (simp add: mult_unit_dvd_iff power_mult_distrib is_unit_power_iff) thus ?thesis by (simp add: multiplicity_def) qed lemma multiplicity_times_unit_right: assumes "is_unit c" shows "multiplicity p (c * x) = multiplicity p x" proof - from assms have "{n. p ^ n dvd c * x} = {n. p ^ n dvd x}" by (subst mult.commute) (simp add: dvd_mult_unit_iff) thus ?thesis by (simp add: multiplicity_def) qed lemma multiplicity_normalize_left [simp]: "multiplicity (normalize p) x = multiplicity p x" proof (cases "p = 0") case [simp]: False have "normalize p = (1 div unit_factor p) * p" by (simp add: unit_div_commute is_unit_unit_factor) also have "multiplicity \ x = multiplicity p x" by (rule multiplicity_times_unit_left) (simp add: is_unit_unit_factor) finally show ?thesis . qed simp_all lemma multiplicity_normalize_right [simp]: "multiplicity p (normalize x) = multiplicity p x" proof (cases "x = 0") case [simp]: False have "normalize x = (1 div unit_factor x) * x" by (simp add: unit_div_commute is_unit_unit_factor) also have "multiplicity p \ = multiplicity p x" by (rule multiplicity_times_unit_right) (simp add: is_unit_unit_factor) finally show ?thesis . qed simp_all lemma multiplicity_prime [simp]: "prime_elem p \ multiplicity p p = 1" by (rule multiplicity_self) auto lemma multiplicity_prime_power [simp]: "prime_elem p \ multiplicity p (p ^ n) = n" by (subst multiplicity_same_power') auto lift_definition prime_factorization :: "'a \ 'a multiset" is "\x p. if prime p then multiplicity p x else 0" proof - fix x :: 'a show "finite {p. 0 < (if prime p then multiplicity p x else 0)}" (is "finite ?A") proof (cases "x = 0") case False from False have "?A \ {p. prime p \ p dvd x}" by (auto simp: multiplicity_gt_zero_iff) moreover from False have "finite {p. prime p \ p dvd x}" by (rule finite_prime_divisors) ultimately show ?thesis by (rule finite_subset) qed simp_all qed abbreviation prime_factors :: "'a \ 'a set" where "prime_factors a \ set_mset (prime_factorization a)" lemma count_prime_factorization_nonprime: "\prime p \ count (prime_factorization x) p = 0" by transfer simp lemma count_prime_factorization_prime: "prime p \ count (prime_factorization x) p = multiplicity p x" by transfer simp lemma count_prime_factorization: "count (prime_factorization x) p = (if prime p then multiplicity p x else 0)" by transfer simp lemma dvd_imp_multiplicity_le: assumes "a dvd b" "b \ 0" shows "multiplicity p a \ multiplicity p b" proof (cases "is_unit p") case False with assms show ?thesis by (intro multiplicity_geI ) (auto intro: dvd_trans[OF multiplicity_dvd' assms(1)]) qed (insert assms, auto simp: multiplicity_unit_left) lemma prime_power_inj: assumes "prime a" "a ^ m = a ^ n" shows "m = n" proof - have "multiplicity a (a ^ m) = multiplicity a (a ^ n)" by (simp only: assms) thus ?thesis using assms by (subst (asm) (1 2) multiplicity_prime_power) simp_all qed lemma prime_power_inj': assumes "prime p" "prime q" assumes "p ^ m = q ^ n" "m > 0" "n > 0" shows "p = q" "m = n" proof - from assms have "p ^ 1 dvd p ^ m" by (intro le_imp_power_dvd) simp also have "p ^ m = q ^ n" by fact finally have "p dvd q ^ n" by simp with assms have "p dvd q" using prime_dvd_power[of p q] by simp with assms show "p = q" by (simp add: primes_dvd_imp_eq) with assms show "m = n" by (simp add: prime_power_inj) qed lemma prime_power_eq_one_iff [simp]: "prime p \ p ^ n = 1 \ n = 0" using prime_power_inj[of p n 0] by auto lemma one_eq_prime_power_iff [simp]: "prime p \ 1 = p ^ n \ n = 0" using prime_power_inj[of p 0 n] by auto lemma prime_power_inj'': assumes "prime p" "prime q" shows "p ^ m = q ^ n \ (m = 0 \ n = 0) \ (p = q \ m = n)" using assms by (cases "m = 0"; cases "n = 0") (auto dest: prime_power_inj'[OF assms]) lemma prime_factorization_0 [simp]: "prime_factorization 0 = {#}" by (simp add: multiset_eq_iff count_prime_factorization) lemma prime_factorization_empty_iff: "prime_factorization x = {#} \ x = 0 \ is_unit x" proof assume *: "prime_factorization x = {#}" { assume x: "x \ 0" "\is_unit x" { fix p assume p: "prime p" have "count (prime_factorization x) p = 0" by (simp add: *) also from p have "count (prime_factorization x) p = multiplicity p x" by (rule count_prime_factorization_prime) also from x p have "\ = 0 \ \p dvd x" by (simp add: multiplicity_eq_zero_iff) finally have "\p dvd x" . } with prime_divisor_exists[OF x] have False by blast } thus "x = 0 \ is_unit x" by blast next assume "x = 0 \ is_unit x" thus "prime_factorization x = {#}" proof assume x: "is_unit x" { fix p assume p: "prime p" from p x have "multiplicity p x = 0" by (subst multiplicity_eq_zero_iff) (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors) } thus ?thesis by (simp add: multiset_eq_iff count_prime_factorization) qed simp_all qed lemma prime_factorization_unit: assumes "is_unit x" shows "prime_factorization x = {#}" proof (rule multiset_eqI) fix p :: 'a show "count (prime_factorization x) p = count {#} p" proof (cases "prime p") case True with assms have "multiplicity p x = 0" by (subst multiplicity_eq_zero_iff) (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors) with True show ?thesis by (simp add: count_prime_factorization_prime) qed (simp_all add: count_prime_factorization_nonprime) qed lemma prime_factorization_1 [simp]: "prime_factorization 1 = {#}" by (simp add: prime_factorization_unit) lemma prime_factorization_times_prime: assumes "x \ 0" "prime p" shows "prime_factorization (p * x) = {#p#} + prime_factorization x" proof (rule multiset_eqI) fix q :: 'a consider "\prime q" | "p = q" | "prime q" "p \ q" by blast thus "count (prime_factorization (p * x)) q = count ({#p#} + prime_factorization x) q" proof cases assume q: "prime q" "p \ q" with assms primes_dvd_imp_eq[of q p] have "\q dvd p" by auto with q assms show ?thesis by (simp add: multiplicity_prime_elem_times_other count_prime_factorization) qed (insert assms, auto simp: count_prime_factorization multiplicity_times_same) qed lemma prod_mset_prime_factorization_weak: assumes "x \ 0" shows "normalize (prod_mset (prime_factorization x)) = normalize x" using assms proof (induction x rule: prime_divisors_induct) case (factor p x) have "normalize (prod_mset (prime_factorization (p * x))) = normalize (p * normalize (prod_mset (prime_factorization x)))" using factor.prems factor.hyps by (simp add: prime_factorization_times_prime) also have "normalize (prod_mset (prime_factorization x)) = normalize x" by (rule factor.IH) (use factor in auto) finally show ?case by simp qed (auto simp: prime_factorization_unit is_unit_normalize) lemma in_prime_factors_iff: "p \ prime_factors x \ x \ 0 \ p dvd x \ prime p" proof - have "p \ prime_factors x \ count (prime_factorization x) p > 0" by simp also have "\ \ x \ 0 \ p dvd x \ prime p" by (subst count_prime_factorization, cases "x = 0") (auto simp: multiplicity_eq_zero_iff multiplicity_gt_zero_iff) finally show ?thesis . qed lemma in_prime_factors_imp_prime [intro]: "p \ prime_factors x \ prime p" by (simp add: in_prime_factors_iff) lemma in_prime_factors_imp_dvd [dest]: "p \ prime_factors x \ p dvd x" by (simp add: in_prime_factors_iff) lemma prime_factorsI: "x \ 0 \ prime p \ p dvd x \ p \ prime_factors x" by (auto simp: in_prime_factors_iff) lemma prime_factors_dvd: "x \ 0 \ prime_factors x = {p. prime p \ p dvd x}" by (auto intro: prime_factorsI) lemma prime_factors_multiplicity: "prime_factors n = {p. prime p \ multiplicity p n > 0}" by (cases "n = 0") (auto simp add: prime_factors_dvd prime_multiplicity_gt_zero_iff) lemma prime_factorization_prime: assumes "prime p" shows "prime_factorization p = {#p#}" proof (rule multiset_eqI) fix q :: 'a consider "\prime q" | "q = p" | "prime q" "q \ p" by blast thus "count (prime_factorization p) q = count {#p#} q" by cases (insert assms, auto dest: primes_dvd_imp_eq simp: count_prime_factorization multiplicity_self multiplicity_eq_zero_iff) qed lemma prime_factorization_prod_mset_primes: assumes "\p. p \# A \ prime p" shows "prime_factorization (prod_mset A) = A" using assms proof (induction A) case (add p A) from add.prems[of 0] have "0 \# A" by auto hence "prod_mset A \ 0" by auto with add show ?case by (simp_all add: mult_ac prime_factorization_times_prime Multiset.union_commute) qed simp_all lemma prime_factorization_cong: "normalize x = normalize y \ prime_factorization x = prime_factorization y" by (simp add: multiset_eq_iff count_prime_factorization multiplicity_normalize_right [of _ x, symmetric] multiplicity_normalize_right [of _ y, symmetric] del: multiplicity_normalize_right) lemma prime_factorization_unique: assumes "x \ 0" "y \ 0" shows "prime_factorization x = prime_factorization y \ normalize x = normalize y" proof assume "prime_factorization x = prime_factorization y" hence "prod_mset (prime_factorization x) = prod_mset (prime_factorization y)" by simp hence "normalize (prod_mset (prime_factorization x)) = normalize (prod_mset (prime_factorization y))" by (simp only: ) with assms show "normalize x = normalize y" by (simp add: prod_mset_prime_factorization_weak) qed (rule prime_factorization_cong) lemma prime_factorization_normalize [simp]: "prime_factorization (normalize x) = prime_factorization x" by (cases "x = 0", simp, subst prime_factorization_unique) auto lemma prime_factorization_eqI_strong: assumes "\p. p \# P \ prime p" "prod_mset P = n" shows "prime_factorization n = P" using prime_factorization_prod_mset_primes[of P] assms by simp lemma prime_factorization_eqI: assumes "\p. p \# P \ prime p" "normalize (prod_mset P) = normalize n" shows "prime_factorization n = P" proof - have "P = prime_factorization (normalize (prod_mset P))" using prime_factorization_prod_mset_primes[of P] assms(1) by simp with assms(2) show ?thesis by simp qed lemma prime_factorization_mult: assumes "x \ 0" "y \ 0" shows "prime_factorization (x * y) = prime_factorization x + prime_factorization y" proof - have "normalize (prod_mset (prime_factorization x) * prod_mset (prime_factorization y)) = normalize (normalize (prod_mset (prime_factorization x)) * normalize (prod_mset (prime_factorization y)))" by (simp only: normalize_mult_normalize_left normalize_mult_normalize_right) also have "\ = normalize (x * y)" by (subst (1 2) prod_mset_prime_factorization_weak) (use assms in auto) finally show ?thesis by (intro prime_factorization_eqI) auto qed lemma prime_factorization_prod: assumes "finite A" "\x. x \ A \ f x \ 0" shows "prime_factorization (prod f A) = (\n\A. prime_factorization (f n))" using assms by (induction A rule: finite_induct) (auto simp: Sup_multiset_empty prime_factorization_mult) lemma prime_elem_multiplicity_mult_distrib: assumes "prime_elem p" "x \ 0" "y \ 0" shows "multiplicity p (x * y) = multiplicity p x + multiplicity p y" proof - have "multiplicity p (x * y) = count (prime_factorization (x * y)) (normalize p)" by (subst count_prime_factorization_prime) (simp_all add: assms) also from assms have "prime_factorization (x * y) = prime_factorization x + prime_factorization y" by (intro prime_factorization_mult) also have "count \ (normalize p) = count (prime_factorization x) (normalize p) + count (prime_factorization y) (normalize p)" by simp also have "\ = multiplicity p x + multiplicity p y" by (subst (1 2) count_prime_factorization_prime) (simp_all add: assms) finally show ?thesis . qed lemma prime_elem_multiplicity_prod_mset_distrib: assumes "prime_elem p" "0 \# A" shows "multiplicity p (prod_mset A) = sum_mset (image_mset (multiplicity p) A)" using assms by (induction A) (auto simp: prime_elem_multiplicity_mult_distrib) lemma prime_elem_multiplicity_power_distrib: assumes "prime_elem p" "x \ 0" shows "multiplicity p (x ^ n) = n * multiplicity p x" using assms prime_elem_multiplicity_prod_mset_distrib [of p "replicate_mset n x"] by simp lemma prime_elem_multiplicity_prod_distrib: assumes "prime_elem p" "0 \ f ` A" "finite A" shows "multiplicity p (prod f A) = (\x\A. multiplicity p (f x))" proof - have "multiplicity p (prod f A) = (\x\#mset_set A. multiplicity p (f x))" using assms by (subst prod_unfold_prod_mset) (simp_all add: prime_elem_multiplicity_prod_mset_distrib sum_unfold_sum_mset multiset.map_comp o_def) also from \finite A\ have "\ = (\x\A. multiplicity p (f x))" by (induction A rule: finite_induct) simp_all finally show ?thesis . qed lemma multiplicity_distinct_prime_power: "prime p \ prime q \ p \ q \ multiplicity p (q ^ n) = 0" by (subst prime_elem_multiplicity_power_distrib) (auto simp: prime_multiplicity_other) lemma prime_factorization_prime_power: "prime p \ prime_factorization (p ^ n) = replicate_mset n p" by (induction n) (simp_all add: prime_factorization_mult prime_factorization_prime Multiset.union_commute) lemma prime_factorization_subset_iff_dvd: assumes [simp]: "x \ 0" "y \ 0" shows "prime_factorization x \# prime_factorization y \ x dvd y" proof - have "x dvd y \ normalize (prod_mset (prime_factorization x)) dvd normalize (prod_mset (prime_factorization y))" using assms by (subst (1 2) prod_mset_prime_factorization_weak) auto also have "\ \ prime_factorization x \# prime_factorization y" by (auto intro!: prod_mset_primes_dvd_imp_subset prod_mset_subset_imp_dvd) finally show ?thesis .. qed lemma prime_factorization_subset_imp_dvd: "x \ 0 \ (prime_factorization x \# prime_factorization y) \ x dvd y" by (cases "y = 0") (simp_all add: prime_factorization_subset_iff_dvd) lemma prime_factorization_divide: assumes "b dvd a" shows "prime_factorization (a div b) = prime_factorization a - prime_factorization b" proof (cases "a = 0") case [simp]: False from assms have [simp]: "b \ 0" by auto have "prime_factorization ((a div b) * b) = prime_factorization (a div b) + prime_factorization b" by (intro prime_factorization_mult) (insert assms, auto elim!: dvdE) with assms show ?thesis by simp qed simp_all lemma zero_not_in_prime_factors [simp]: "0 \ prime_factors x" by (auto dest: in_prime_factors_imp_prime) lemma prime_prime_factors: "prime p \ prime_factors p = {p}" by (drule prime_factorization_prime) simp lemma prime_factors_product: "x \ 0 \ y \ 0 \ prime_factors (x * y) = prime_factors x \ prime_factors y" by (simp add: prime_factorization_mult) lemma dvd_prime_factors [intro]: "y \ 0 \ x dvd y \ prime_factors x \ prime_factors y" by (intro set_mset_mono, subst prime_factorization_subset_iff_dvd) auto (* RENAMED multiplicity_dvd *) lemma multiplicity_le_imp_dvd: assumes "x \ 0" "\p. prime p \ multiplicity p x \ multiplicity p y" shows "x dvd y" proof (cases "y = 0") case False from assms this have "prime_factorization x \# prime_factorization y" by (intro mset_subset_eqI) (auto simp: count_prime_factorization) with assms False show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) qed auto lemma dvd_multiplicity_eq: "x \ 0 \ y \ 0 \ x dvd y \ (\p. multiplicity p x \ multiplicity p y)" by (auto intro: dvd_imp_multiplicity_le multiplicity_le_imp_dvd) lemma multiplicity_eq_imp_eq: assumes "x \ 0" "y \ 0" assumes "\p. prime p \ multiplicity p x = multiplicity p y" shows "normalize x = normalize y" using assms by (intro associatedI multiplicity_le_imp_dvd) simp_all lemma prime_factorization_unique': assumes "\p \# M. prime p" "\p \# N. prime p" "(\i \# M. i) = (\i \# N. i)" shows "M = N" proof - have "prime_factorization (\i \# M. i) = prime_factorization (\i \# N. i)" by (simp only: assms) also from assms have "prime_factorization (\i \# M. i) = M" by (subst prime_factorization_prod_mset_primes) simp_all also from assms have "prime_factorization (\i \# N. i) = N" by (subst prime_factorization_prod_mset_primes) simp_all finally show ?thesis . qed lemma prime_factorization_unique'': assumes "\p \# M. prime p" "\p \# N. prime p" "normalize (\i \# M. i) = normalize (\i \# N. i)" shows "M = N" proof - have "prime_factorization (normalize (\i \# M. i)) = prime_factorization (normalize (\i \# N. i))" by (simp only: assms) also from assms have "prime_factorization (normalize (\i \# M. i)) = M" by (subst prime_factorization_normalize, subst prime_factorization_prod_mset_primes) simp_all also from assms have "prime_factorization (normalize (\i \# N. i)) = N" by (subst prime_factorization_normalize, subst prime_factorization_prod_mset_primes) simp_all finally show ?thesis . qed lemma multiplicity_cong: "(\r. p ^ r dvd a \ p ^ r dvd b) \ multiplicity p a = multiplicity p b" by (simp add: multiplicity_def) lemma not_dvd_imp_multiplicity_0: assumes "\p dvd x" shows "multiplicity p x = 0" proof - from assms have "multiplicity p x < 1" by (intro multiplicity_lessI) auto thus ?thesis by simp qed lemma multiplicity_zero_left [simp]: "multiplicity 0 x = 0" by (cases "x = 0") (auto intro: not_dvd_imp_multiplicity_0) lemma inj_on_Prod_primes: assumes "\P p. P \ A \ p \ P \ prime p" assumes "\P. P \ A \ finite P" shows "inj_on Prod A" proof (rule inj_onI) fix P Q assume PQ: "P \ A" "Q \ A" "\P = \Q" with prime_factorization_unique'[of "mset_set P" "mset_set Q"] assms[of P] assms[of Q] have "mset_set P = mset_set Q" by (auto simp: prod_unfold_prod_mset) with assms[of P] assms[of Q] PQ show "P = Q" by simp qed lemma divides_primepow_weak: assumes "prime p" and "a dvd p ^ n" obtains m where "m \ n" and "normalize a = normalize (p ^ m)" proof - from assms have "a \ 0" by auto with assms have "normalize (prod_mset (prime_factorization a)) dvd normalize (prod_mset (prime_factorization (p ^ n)))" by (subst (1 2) prod_mset_prime_factorization_weak) auto then have "prime_factorization a \# prime_factorization (p ^ n)" by (simp add: in_prime_factors_imp_prime prod_mset_dvd_prod_mset_primes_iff) with assms have "prime_factorization a \# replicate_mset n p" by (simp add: prime_factorization_prime_power) then obtain m where "m \ n" and "prime_factorization a = replicate_mset m p" by (rule msubseteq_replicate_msetE) then have *: "normalize (prod_mset (prime_factorization a)) = normalize (prod_mset (replicate_mset m p))" by metis also have "normalize (prod_mset (prime_factorization a)) = normalize a" using \a \ 0\ by (simp add: prod_mset_prime_factorization_weak) also have "prod_mset (replicate_mset m p) = p ^ m" by simp finally show ?thesis using \m \ n\ by (intro that[of m]) qed lemma divide_out_primepow_ex: assumes "n \ 0" "\p\prime_factors n. P p" obtains p k n' where "P p" "prime p" "p dvd n" "\p dvd n'" "k > 0" "n = p ^ k * n'" proof - from assms obtain p where p: "P p" "prime p" "p dvd n" by auto define k where "k = multiplicity p n" define n' where "n' = n div p ^ k" have n': "n = p ^ k * n'" "\p dvd n'" using assms p multiplicity_decompose[of n p] by (auto simp: n'_def k_def multiplicity_dvd) from n' p have "k > 0" by (intro Nat.gr0I) auto with n' p that[of p n' k] show ?thesis by auto qed lemma divide_out_primepow: assumes "n \ 0" "\is_unit n" obtains p k n' where "prime p" "p dvd n" "\p dvd n'" "k > 0" "n = p ^ k * n'" using divide_out_primepow_ex[OF assms(1), of "\_. True"] prime_divisor_exists[OF assms] assms prime_factorsI by metis subsection \GCD and LCM computation with unique factorizations\ definition "gcd_factorial a b = (if a = 0 then normalize b else if b = 0 then normalize a else normalize (prod_mset (prime_factorization a \# prime_factorization b)))" definition "lcm_factorial a b = (if a = 0 \ b = 0 then 0 else normalize (prod_mset (prime_factorization a \# prime_factorization b)))" definition "Gcd_factorial A = (if A \ {0} then 0 else normalize (prod_mset (Inf (prime_factorization ` (A - {0})))))" definition "Lcm_factorial A = (if A = {} then 1 else if 0 \ A \ subset_mset.bdd_above (prime_factorization ` (A - {0})) then normalize (prod_mset (Sup (prime_factorization ` A))) else 0)" lemma prime_factorization_gcd_factorial: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (gcd_factorial a b) = prime_factorization a \# prime_factorization b" proof - have "prime_factorization (gcd_factorial a b) = prime_factorization (prod_mset (prime_factorization a \# prime_factorization b))" by (simp add: gcd_factorial_def) also have "\ = prime_factorization a \# prime_factorization b" by (subst prime_factorization_prod_mset_primes) auto finally show ?thesis . qed lemma prime_factorization_lcm_factorial: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (lcm_factorial a b) = prime_factorization a \# prime_factorization b" proof - have "prime_factorization (lcm_factorial a b) = prime_factorization (prod_mset (prime_factorization a \# prime_factorization b))" by (simp add: lcm_factorial_def) also have "\ = prime_factorization a \# prime_factorization b" by (subst prime_factorization_prod_mset_primes) auto finally show ?thesis . qed lemma prime_factorization_Gcd_factorial: assumes "\A \ {0}" shows "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))" proof - from assms obtain x where x: "x \ A - {0}" by auto hence "Inf (prime_factorization ` (A - {0})) \# prime_factorization x" by (intro subset_mset.cInf_lower) simp_all hence "\y. y \# Inf (prime_factorization ` (A - {0})) \ y \ prime_factors x" by (auto dest: mset_subset_eqD) with in_prime_factors_imp_prime[of _ x] have "\p. p \# Inf (prime_factorization ` (A - {0})) \ prime p" by blast with assms show ?thesis by (simp add: Gcd_factorial_def prime_factorization_prod_mset_primes) qed lemma prime_factorization_Lcm_factorial: assumes "0 \ A" "subset_mset.bdd_above (prime_factorization ` A)" shows "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)" proof (cases "A = {}") case True hence "prime_factorization ` A = {}" by auto also have "Sup \ = {#}" by (simp add: Sup_multiset_empty) finally show ?thesis by (simp add: Lcm_factorial_def) next case False have "\y. y \# Sup (prime_factorization ` A) \ prime y" by (auto simp: in_Sup_multiset_iff assms) with assms False show ?thesis by (simp add: Lcm_factorial_def prime_factorization_prod_mset_primes) qed lemma gcd_factorial_commute: "gcd_factorial a b = gcd_factorial b a" by (simp add: gcd_factorial_def multiset_inter_commute) lemma gcd_factorial_dvd1: "gcd_factorial a b dvd a" proof (cases "a = 0 \ b = 0") case False hence "gcd_factorial a b \ 0" by (auto simp: gcd_factorial_def) with False show ?thesis by (subst prime_factorization_subset_iff_dvd [symmetric]) (auto simp: prime_factorization_gcd_factorial) qed (auto simp: gcd_factorial_def) lemma gcd_factorial_dvd2: "gcd_factorial a b dvd b" by (subst gcd_factorial_commute) (rule gcd_factorial_dvd1) lemma normalize_gcd_factorial [simp]: "normalize (gcd_factorial a b) = gcd_factorial a b" by (simp add: gcd_factorial_def) lemma normalize_lcm_factorial [simp]: "normalize (lcm_factorial a b) = lcm_factorial a b" by (simp add: lcm_factorial_def) lemma gcd_factorial_greatest: "c dvd gcd_factorial a b" if "c dvd a" "c dvd b" for a b c proof (cases "a = 0 \ b = 0") case False with that have [simp]: "c \ 0" by auto let ?p = "prime_factorization" from that False have "?p c \# ?p a" "?p c \# ?p b" by (simp_all add: prime_factorization_subset_iff_dvd) hence "prime_factorization c \# prime_factorization (prod_mset (prime_factorization a \# prime_factorization b))" using False by (subst prime_factorization_prod_mset_primes) auto with False show ?thesis by (auto simp: gcd_factorial_def prime_factorization_subset_iff_dvd [symmetric]) qed (auto simp: gcd_factorial_def that) lemma lcm_factorial_gcd_factorial: "lcm_factorial a b = normalize (a * b div gcd_factorial a b)" for a b proof (cases "a = 0 \ b = 0") case False let ?p = "prime_factorization" have 1: "normalize x * normalize y dvd z \ x * y dvd z" for x y z :: 'a proof - have "normalize (normalize x * normalize y) dvd z \ x * y dvd z" unfolding normalize_mult_normalize_left normalize_mult_normalize_right by simp thus ?thesis unfolding normalize_dvd_iff by simp qed have "?p (a * b) = (?p a \# ?p b) + (?p a \# ?p b)" using False by (subst prime_factorization_mult) (auto intro!: multiset_eqI) hence "normalize (prod_mset (?p (a * b))) = normalize (prod_mset ((?p a \# ?p b) + (?p a \# ?p b)))" by (simp only:) hence *: "normalize (a * b) = normalize (lcm_factorial a b * gcd_factorial a b)" using False by (subst (asm) prod_mset_prime_factorization_weak) (auto simp: lcm_factorial_def gcd_factorial_def) have [simp]: "gcd_factorial a b dvd a * b" "lcm_factorial a b dvd a * b" using associatedD2[OF *] by auto from False have [simp]: "gcd_factorial a b \ 0" "lcm_factorial a b \ 0" by (auto simp: gcd_factorial_def lcm_factorial_def) show ?thesis by (rule associated_eqI) (use * in \auto simp: dvd_div_iff_mult div_dvd_iff_mult dest: associatedD1 associatedD2\) qed (auto simp: lcm_factorial_def) lemma normalize_Gcd_factorial: "normalize (Gcd_factorial A) = Gcd_factorial A" by (simp add: Gcd_factorial_def) lemma Gcd_factorial_eq_0_iff: "Gcd_factorial A = 0 \ A \ {0}" by (auto simp: Gcd_factorial_def in_Inf_multiset_iff split: if_splits) lemma Gcd_factorial_dvd: assumes "x \ A" shows "Gcd_factorial A dvd x" proof (cases "x = 0") case False with assms have "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))" by (intro prime_factorization_Gcd_factorial) auto also from False assms have "\ \# prime_factorization x" by (intro subset_mset.cInf_lower) auto finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert assms False, auto simp: Gcd_factorial_eq_0_iff) qed simp_all lemma Gcd_factorial_greatest: assumes "\y. y \ A \ x dvd y" shows "x dvd Gcd_factorial A" proof (cases "A \ {0}") case False from False obtain y where "y \ A" "y \ 0" by auto with assms[of y] have nz: "x \ 0" by auto from nz assms have "prime_factorization x \# prime_factorization y" if "y \ A - {0}" for y using that by (subst prime_factorization_subset_iff_dvd) auto with False have "prime_factorization x \# Inf (prime_factorization ` (A - {0}))" by (intro subset_mset.cInf_greatest) auto also from False have "\ = prime_factorization (Gcd_factorial A)" by (rule prime_factorization_Gcd_factorial [symmetric]) finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert nz False, auto simp: Gcd_factorial_eq_0_iff) qed (simp_all add: Gcd_factorial_def) lemma normalize_Lcm_factorial: "normalize (Lcm_factorial A) = Lcm_factorial A" by (simp add: Lcm_factorial_def) lemma Lcm_factorial_eq_0_iff: "Lcm_factorial A = 0 \ 0 \ A \ \subset_mset.bdd_above (prime_factorization ` A)" by (auto simp: Lcm_factorial_def in_Sup_multiset_iff) lemma dvd_Lcm_factorial: assumes "x \ A" shows "x dvd Lcm_factorial A" proof (cases "0 \ A \ subset_mset.bdd_above (prime_factorization ` A)") case True with assms have [simp]: "0 \ A" "x \ 0" "A \ {}" by auto from assms True have "prime_factorization x \# Sup (prime_factorization ` A)" by (intro subset_mset.cSup_upper) auto also have "\ = prime_factorization (Lcm_factorial A)" by (rule prime_factorization_Lcm_factorial [symmetric]) (insert True, simp_all) finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert True, auto simp: Lcm_factorial_eq_0_iff) qed (insert assms, auto simp: Lcm_factorial_def) lemma Lcm_factorial_least: assumes "\y. y \ A \ y dvd x" shows "Lcm_factorial A dvd x" proof - consider "A = {}" | "0 \ A" | "x = 0" | "A \ {}" "0 \ A" "x \ 0" by blast thus ?thesis proof cases assume *: "A \ {}" "0 \ A" "x \ 0" hence nz: "x \ 0" if "x \ A" for x using that by auto from * have bdd: "subset_mset.bdd_above (prime_factorization ` A)" by (intro subset_mset.bdd_aboveI[of _ "prime_factorization x"]) (auto simp: prime_factorization_subset_iff_dvd nz dest: assms) have "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)" by (rule prime_factorization_Lcm_factorial) fact+ also from * have "\ \# prime_factorization x" by (intro subset_mset.cSup_least) (auto simp: prime_factorization_subset_iff_dvd nz dest: assms) finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert * bdd, auto simp: Lcm_factorial_eq_0_iff) qed (auto simp: Lcm_factorial_def dest: assms) qed lemmas gcd_lcm_factorial = gcd_factorial_dvd1 gcd_factorial_dvd2 gcd_factorial_greatest normalize_gcd_factorial lcm_factorial_gcd_factorial normalize_Gcd_factorial Gcd_factorial_dvd Gcd_factorial_greatest normalize_Lcm_factorial dvd_Lcm_factorial Lcm_factorial_least end class factorial_semiring_gcd = factorial_semiring + gcd + Gcd + assumes gcd_eq_gcd_factorial: "gcd a b = gcd_factorial a b" and lcm_eq_lcm_factorial: "lcm a b = lcm_factorial a b" and Gcd_eq_Gcd_factorial: "Gcd A = Gcd_factorial A" and Lcm_eq_Lcm_factorial: "Lcm A = Lcm_factorial A" begin lemma prime_factorization_gcd: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (gcd a b) = prime_factorization a \# prime_factorization b" by (simp add: gcd_eq_gcd_factorial prime_factorization_gcd_factorial) lemma prime_factorization_lcm: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (lcm a b) = prime_factorization a \# prime_factorization b" by (simp add: lcm_eq_lcm_factorial prime_factorization_lcm_factorial) lemma prime_factorization_Gcd: assumes "Gcd A \ 0" shows "prime_factorization (Gcd A) = Inf (prime_factorization ` (A - {0}))" using assms by (simp add: prime_factorization_Gcd_factorial Gcd_eq_Gcd_factorial Gcd_factorial_eq_0_iff) lemma prime_factorization_Lcm: assumes "Lcm A \ 0" shows "prime_factorization (Lcm A) = Sup (prime_factorization ` A)" using assms by (simp add: prime_factorization_Lcm_factorial Lcm_eq_Lcm_factorial Lcm_factorial_eq_0_iff) lemma prime_factors_gcd [simp]: "a \ 0 \ b \ 0 \ prime_factors (gcd a b) = prime_factors a \ prime_factors b" by (subst prime_factorization_gcd) auto lemma prime_factors_lcm [simp]: "a \ 0 \ b \ 0 \ prime_factors (lcm a b) = prime_factors a \ prime_factors b" by (subst prime_factorization_lcm) auto subclass semiring_gcd by (standard, unfold gcd_eq_gcd_factorial lcm_eq_lcm_factorial) (rule gcd_lcm_factorial; assumption)+ subclass semiring_Gcd by (standard, unfold Gcd_eq_Gcd_factorial Lcm_eq_Lcm_factorial) (rule gcd_lcm_factorial; assumption)+ lemma assumes "x \ 0" "y \ 0" shows gcd_eq_factorial': "gcd x y = normalize (\p \ prime_factors x \ prime_factors y. p ^ min (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs1") and lcm_eq_factorial': "lcm x y = normalize (\p \ prime_factors x \ prime_factors y. p ^ max (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs2") proof - have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial) also have "\ = ?rhs1" by (auto simp: gcd_factorial_def assms prod_mset_multiplicity count_prime_factorization_prime intro!: arg_cong[of _ _ normalize] dest: in_prime_factors_imp_prime intro!: prod.cong) finally show "gcd x y = ?rhs1" . have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial) also have "\ = ?rhs2" by (auto simp: lcm_factorial_def assms prod_mset_multiplicity count_prime_factorization_prime intro!: arg_cong[of _ _ normalize] dest: in_prime_factors_imp_prime intro!: prod.cong) finally show "lcm x y = ?rhs2" . qed lemma assumes "x \ 0" "y \ 0" "prime p" shows multiplicity_gcd: "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)" and multiplicity_lcm: "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)" proof - have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial) also from assms have "multiplicity p \ = min (multiplicity p x) (multiplicity p y)" by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_gcd_factorial) finally show "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)" . have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial) also from assms have "multiplicity p \ = max (multiplicity p x) (multiplicity p y)" by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_lcm_factorial) finally show "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)" . qed lemma gcd_lcm_distrib: "gcd x (lcm y z) = lcm (gcd x y) (gcd x z)" proof (cases "x = 0 \ y = 0 \ z = 0") case True thus ?thesis by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd) next case False hence "normalize (gcd x (lcm y z)) = normalize (lcm (gcd x y) (gcd x z))" by (intro associatedI prime_factorization_subset_imp_dvd) (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm subset_mset.inf_sup_distrib1) thus ?thesis by simp qed lemma lcm_gcd_distrib: "lcm x (gcd y z) = gcd (lcm x y) (lcm x z)" proof (cases "x = 0 \ y = 0 \ z = 0") case True thus ?thesis by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd) next case False hence "normalize (lcm x (gcd y z)) = normalize (gcd (lcm x y) (lcm x z))" by (intro associatedI prime_factorization_subset_imp_dvd) (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm subset_mset.sup_inf_distrib1) thus ?thesis by simp qed end class factorial_ring_gcd = factorial_semiring_gcd + idom begin subclass ring_gcd .. subclass idom_divide .. end class factorial_semiring_multiplicative = factorial_semiring + normalization_semidom_multiplicative begin lemma normalize_prod_mset_primes: "(\p. p \# A \ prime p) \ normalize (prod_mset A) = prod_mset A" proof (induction A) case (add p A) hence "prime p" by simp hence "normalize p = p" by simp with add show ?case by (simp add: normalize_mult) qed simp_all lemma prod_mset_prime_factorization: assumes "x \ 0" shows "prod_mset (prime_factorization x) = normalize x" using assms by (induction x rule: prime_divisors_induct) (simp_all add: prime_factorization_unit prime_factorization_times_prime is_unit_normalize normalize_mult) lemma prime_decomposition: "unit_factor x * prod_mset (prime_factorization x) = x" by (cases "x = 0") (simp_all add: prod_mset_prime_factorization) lemma prod_prime_factors: assumes "x \ 0" shows "(\p \ prime_factors x. p ^ multiplicity p x) = normalize x" proof - have "normalize x = prod_mset (prime_factorization x)" by (simp add: prod_mset_prime_factorization assms) also have "\ = (\p \ prime_factors x. p ^ count (prime_factorization x) p)" by (subst prod_mset_multiplicity) simp_all also have "\ = (\p \ prime_factors x. p ^ multiplicity p x)" by (intro prod.cong) (simp_all add: assms count_prime_factorization_prime in_prime_factors_imp_prime) finally show ?thesis .. qed lemma prime_factorization_unique'': assumes S_eq: "S = {p. 0 < f p}" and "finite S" and S: "\p\S. prime p" "normalize n = (\p\S. p ^ f p)" shows "S = prime_factors n \ (\p. prime p \ f p = multiplicity p n)" proof define A where "A = Abs_multiset f" from \finite S\ S(1) have "(\p\S. p ^ f p) \ 0" by auto with S(2) have nz: "n \ 0" by auto from S_eq \finite S\ have count_A: "count A = f" unfolding A_def by (subst multiset.Abs_multiset_inverse) simp_all from S_eq count_A have set_mset_A: "set_mset A = S" by (simp only: set_mset_def) from S(2) have "normalize n = (\p\S. p ^ f p)" . also have "\ = prod_mset A" by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A) also from nz have "normalize n = prod_mset (prime_factorization n)" by (simp add: prod_mset_prime_factorization) finally have "prime_factorization (prod_mset A) = prime_factorization (prod_mset (prime_factorization n))" by simp also from S(1) have "prime_factorization (prod_mset A) = A" by (intro prime_factorization_prod_mset_primes) (auto simp: set_mset_A) also have "prime_factorization (prod_mset (prime_factorization n)) = prime_factorization n" by (intro prime_factorization_prod_mset_primes) auto finally show "S = prime_factors n" by (simp add: set_mset_A [symmetric]) show "(\p. prime p \ f p = multiplicity p n)" proof safe fix p :: 'a assume p: "prime p" have "multiplicity p n = multiplicity p (normalize n)" by simp also have "normalize n = prod_mset A" by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A S) also from p set_mset_A S(1) have "multiplicity p \ = sum_mset (image_mset (multiplicity p) A)" by (intro prime_elem_multiplicity_prod_mset_distrib) auto also from S(1) p have "image_mset (multiplicity p) A = image_mset (\q. if p = q then 1 else 0) A" by (intro image_mset_cong) (auto simp: set_mset_A multiplicity_self prime_multiplicity_other) also have "sum_mset \ = f p" by (simp add: semiring_1_class.sum_mset_delta' count_A) finally show "f p = multiplicity p n" .. qed qed lemma divides_primepow: assumes "prime p" and "a dvd p ^ n" obtains m where "m \ n" and "normalize a = p ^ m" using divides_primepow_weak[OF assms] that assms by (auto simp add: normalize_power) lemma Ex_other_prime_factor: assumes "n \ 0" and "\(\k. normalize n = p ^ k)" "prime p" shows "\q\prime_factors n. q \ p" proof (rule ccontr) assume *: "\(\q\prime_factors n. q \ p)" have "normalize n = (\p\prime_factors n. p ^ multiplicity p n)" using assms(1) by (intro prod_prime_factors [symmetric]) auto also from * have "\ = (\p\{p}. p ^ multiplicity p n)" using assms(3) by (intro prod.mono_neutral_left) (auto simp: prime_factors_multiplicity) finally have "normalize n = p ^ multiplicity p n" by auto with assms show False by auto qed text \Now a string of results due to Jakub Kądziołka\ lemma multiplicity_dvd_iff_dvd: assumes "x \ 0" shows "p^k dvd x \ p^k dvd p^multiplicity p x" proof (cases "is_unit p") case True then have "is_unit (p^k)" using is_unit_power_iff by simp hence "p^k dvd x" by auto - moreover from `is_unit p` have "p^k dvd p^multiplicity p x" + moreover from \is_unit p\ have "p^k dvd p^multiplicity p x" using multiplicity_unit_left is_unit_power_iff by simp ultimately show ?thesis by simp next case False show ?thesis proof (cases "p = 0") case True then have "p^multiplicity p x = 1" by simp moreover have "p^k dvd x \ k = 0" proof (rule ccontr) assume "p^k dvd x" and "k \ 0" - with `p = 0` have "p^k = 0" by auto - with `p^k dvd x` have "0 dvd x" by auto + with \p = 0\ have "p^k = 0" by auto + with \p^k dvd x\ have "0 dvd x" by auto hence "x = 0" by auto - with `x \ 0` show False by auto + with \x \ 0\ show False by auto qed ultimately show ?thesis - by (auto simp add: is_unit_power_iff `\ is_unit p`) + by (auto simp add: is_unit_power_iff \\ is_unit p\) next case False - with `x \ 0` `\ is_unit p` show ?thesis + with \x \ 0\ \\ is_unit p\ show ?thesis by (simp add: power_dvd_iff_le_multiplicity dvd_power_iff multiplicity_same_power) qed qed lemma multiplicity_decomposeI: assumes "x = p^k * x'" and "\ p dvd x'" and "p \ 0" shows "multiplicity p x = k" using assms local.multiplicity_eqI local.power_Suc2 by force lemma multiplicity_sum_lt: assumes "multiplicity p a < multiplicity p b" "a \ 0" "b \ 0" shows "multiplicity p (a + b) = multiplicity p a" proof - let ?vp = "multiplicity p" have unit: "\ is_unit p" proof assume "is_unit p" then have "?vp a = 0" and "?vp b = 0" using multiplicity_unit_left by auto with assms show False by auto qed from multiplicity_decompose' obtain a' where a': "a = p^?vp a * a'" "\ p dvd a'" using unit assms by metis from multiplicity_decompose' obtain b' where b': "b = p^?vp b * b'" using unit assms by metis show "?vp (a + b) = ?vp a" proof (rule multiplicity_decomposeI) let ?k = "?vp b - ?vp a" from assms have k: "?k > 0" by simp with b' have "b = p^?vp a * p^?k * b'" by (simp flip: power_add) with a' show *: "a + b = p^?vp a * (a' + p^?k * b')" by (simp add: ac_simps distrib_left) moreover show "\ p dvd a' + p^?k * b'" using a' k dvd_add_left_iff by auto show "p \ 0" using assms by auto qed qed corollary multiplicity_sum_min: assumes "multiplicity p a \ multiplicity p b" "a \ 0" "b \ 0" shows "multiplicity p (a + b) = min (multiplicity p a) (multiplicity p b)" proof - let ?vp = "multiplicity p" from assms have "?vp a < ?vp b \ ?vp a > ?vp b" by auto then show ?thesis by (metis assms multiplicity_sum_lt min.commute add_commute min.strict_order_iff) qed end end diff --git a/src/HOL/Equiv_Relations.thy b/src/HOL/Equiv_Relations.thy --- a/src/HOL/Equiv_Relations.thy +++ b/src/HOL/Equiv_Relations.thy @@ -1,642 +1,642 @@ (* Title: HOL/Equiv_Relations.thy Author: Lawrence C Paulson, 1996 Cambridge University Computer Laboratory *) section \Equivalence Relations in Higher-Order Set Theory\ theory Equiv_Relations imports Groups_Big begin subsection \Equivalence relations -- set version\ definition equiv :: "'a set \ ('a \ 'a) set \ bool" where "equiv A r \ refl_on A r \ sym r \ trans r" lemma equivI: "refl_on A r \ sym r \ trans r \ equiv A r" by (simp add: equiv_def) lemma equivE: assumes "equiv A r" obtains "refl_on A r" and "sym r" and "trans r" using assms by (simp add: equiv_def) text \ Suppes, Theorem 70: \r\ is an equiv relation iff \r\ O r = r\. First half: \equiv A r \ r\ O r = r\. \ lemma sym_trans_comp_subset: "sym r \ trans r \ r\ O r \ r" unfolding trans_def sym_def converse_unfold by blast lemma refl_on_comp_subset: "refl_on A r \ r \ r\ O r" unfolding refl_on_def by blast lemma equiv_comp_eq: "equiv A r \ r\ O r = r" unfolding equiv_def by (iprover intro: sym_trans_comp_subset refl_on_comp_subset equalityI) text \Second half.\ lemma comp_equivI: assumes "r\ O r = r" "Domain r = A" shows "equiv A r" proof - have *: "\x y. (x, y) \ r \ (y, x) \ r" using assms by blast show ?thesis unfolding equiv_def refl_on_def sym_def trans_def using assms by (auto intro: *) qed subsection \Equivalence classes\ lemma equiv_class_subset: "equiv A r \ (a, b) \ r \ r``{a} \ r``{b}" \ \lemma for the next result\ unfolding equiv_def trans_def sym_def by blast theorem equiv_class_eq: "equiv A r \ (a, b) \ r \ r``{a} = r``{b}" by (intro equalityI equiv_class_subset; force simp add: equiv_def sym_def) lemma equiv_class_self: "equiv A r \ a \ A \ a \ r``{a}" unfolding equiv_def refl_on_def by blast lemma subset_equiv_class: "equiv A r \ r``{b} \ r``{a} \ b \ A \ (a, b) \ r" \ \lemma for the next result\ unfolding equiv_def refl_on_def by blast lemma eq_equiv_class: "r``{a} = r``{b} \ equiv A r \ b \ A \ (a, b) \ r" by (iprover intro: equalityD2 subset_equiv_class) lemma equiv_class_nondisjoint: "equiv A r \ x \ (r``{a} \ r``{b}) \ (a, b) \ r" unfolding equiv_def trans_def sym_def by blast lemma equiv_type: "equiv A r \ r \ A \ A" unfolding equiv_def refl_on_def by blast lemma equiv_class_eq_iff: "equiv A r \ (x, y) \ r \ r``{x} = r``{y} \ x \ A \ y \ A" by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) lemma eq_equiv_class_iff: "equiv A r \ x \ A \ y \ A \ r``{x} = r``{y} \ (x, y) \ r" by (blast intro!: equiv_class_eq dest: eq_equiv_class equiv_type) lemma disjnt_equiv_class: "equiv A r \ disjnt (r``{a}) (r``{b}) \ (a, b) \ r" by (auto dest: equiv_class_self simp: equiv_class_eq_iff disjnt_def) subsection \Quotients\ definition quotient :: "'a set \ ('a \ 'a) set \ 'a set set" (infixl "'/'/" 90) where "A//r = (\x \ A. {r``{x}})" \ \set of equiv classes\ lemma quotientI: "x \ A \ r``{x} \ A//r" unfolding quotient_def by blast lemma quotientE: "X \ A//r \ (\x. X = r``{x} \ x \ A \ P) \ P" unfolding quotient_def by blast lemma Union_quotient: "equiv A r \ \(A//r) = A" unfolding equiv_def refl_on_def quotient_def by blast lemma quotient_disj: "equiv A r \ X \ A//r \ Y \ A//r \ X = Y \ X \ Y = {}" unfolding quotient_def equiv_def trans_def sym_def by blast lemma quotient_eqI: assumes "equiv A r" "X \ A//r" "Y \ A//r" and xy: "x \ X" "y \ Y" "(x, y) \ r" shows "X = Y" proof - obtain a b where "a \ A" and a: "X = r `` {a}" and "b \ A" and b: "Y = r `` {b}" using assms by (auto elim!: quotientE) then have "(a,b) \ r" using xy \equiv A r\ unfolding equiv_def sym_def trans_def by blast then show ?thesis unfolding a b by (rule equiv_class_eq [OF \equiv A r\]) qed lemma quotient_eq_iff: assumes "equiv A r" "X \ A//r" "Y \ A//r" and xy: "x \ X" "y \ Y" shows "X = Y \ (x, y) \ r" proof assume L: "X = Y" with assms show "(x, y) \ r" unfolding equiv_def sym_def trans_def by (blast elim!: quotientE) next assume \
: "(x, y) \ r" show "X = Y" by (rule quotient_eqI) (use \
assms in \blast+\) qed lemma eq_equiv_class_iff2: "equiv A r \ x \ A \ y \ A \ {x}//r = {y}//r \ (x, y) \ r" by (simp add: quotient_def eq_equiv_class_iff) lemma quotient_empty [simp]: "{}//r = {}" by (simp add: quotient_def) lemma quotient_is_empty [iff]: "A//r = {} \ A = {}" by (simp add: quotient_def) lemma quotient_is_empty2 [iff]: "{} = A//r \ A = {}" by (simp add: quotient_def) lemma singleton_quotient: "{x}//r = {r `` {x}}" by (simp add: quotient_def) lemma quotient_diff1: "inj_on (\a. {a}//r) A \ a \ A \ (A - {a})//r = A//r - {a}//r" unfolding quotient_def inj_on_def by blast subsection \Refinement of one equivalence relation WRT another\ lemma refines_equiv_class_eq: "R \ S \ equiv A R \ equiv A S \ R``(S``{a}) = S``{a}" by (auto simp: equiv_class_eq_iff) lemma refines_equiv_class_eq2: "R \ S \ equiv A R \ equiv A S \ S``(R``{a}) = S``{a}" by (auto simp: equiv_class_eq_iff) lemma refines_equiv_image_eq: "R \ S \ equiv A R \ equiv A S \ (\X. S``X) ` (A//R) = A//S" by (auto simp: quotient_def image_UN refines_equiv_class_eq2) lemma finite_refines_finite: "finite (A//R) \ R \ S \ equiv A R \ equiv A S \ finite (A//S)" by (erule finite_surj [where f = "\X. S``X"]) (simp add: refines_equiv_image_eq) lemma finite_refines_card_le: "finite (A//R) \ R \ S \ equiv A R \ equiv A S \ card (A//S) \ card (A//R)" by (subst refines_equiv_image_eq [of R S A, symmetric]) (auto simp: card_image_le [where f = "\X. S``X"]) subsection \Defining unary operations upon equivalence classes\ text \A congruence-preserving function.\ definition congruent :: "('a \ 'a) set \ ('a \ 'b) \ bool" where "congruent r f \ (\(y, z) \ r. f y = f z)" lemma congruentI: "(\y z. (y, z) \ r \ f y = f z) \ congruent r f" by (auto simp add: congruent_def) lemma congruentD: "congruent r f \ (y, z) \ r \ f y = f z" by (auto simp add: congruent_def) abbreviation RESPECTS :: "('a \ 'b) \ ('a \ 'a) set \ bool" (infixr "respects" 80) where "f respects r \ congruent r f" lemma UN_constant_eq: "a \ A \ \y \ A. f y = c \ (\y \ A. f y) = c" \ \lemma required to prove \UN_equiv_class\\ by auto lemma UN_equiv_class: assumes "equiv A r" "f respects r" "a \ A" shows "(\x \ r``{a}. f x) = f a" \ \Conversion rule\ proof - have \
: "\x\r `` {a}. f x = f a" using assms unfolding equiv_def congruent_def sym_def by blast show ?thesis by (iprover intro: assms UN_constant_eq [OF equiv_class_self \
]) qed lemma UN_equiv_class_type: assumes r: "equiv A r" "f respects r" and X: "X \ A//r" and AB: "\x. x \ A \ f x \ B" shows "(\x \ X. f x) \ B" using assms unfolding quotient_def by (auto simp: UN_equiv_class [OF r]) text \ Sufficient conditions for injectiveness. Could weaken premises! major premise could be an inclusion; \bcong\ could be \\y. y \ A \ f y \ B\. \ lemma UN_equiv_class_inject: assumes "equiv A r" "f respects r" and eq: "(\x \ X. f x) = (\y \ Y. f y)" and X: "X \ A//r" and Y: "Y \ A//r" and fr: "\x y. x \ A \ y \ A \ f x = f y \ (x, y) \ r" shows "X = Y" proof - obtain a b where "a \ A" and a: "X = r `` {a}" and "b \ A" and b: "Y = r `` {b}" using assms by (auto elim!: quotientE) then have "\ (f ` r `` {a}) = f a" "\ (f ` r `` {b}) = f b" by (iprover intro: UN_equiv_class [OF \equiv A r\] assms)+ then have "f a = f b" using eq unfolding a b by (iprover intro: trans sym) then have "(a,b) \ r" using fr \a \ A\ \b \ A\ by blast then show ?thesis unfolding a b by (rule equiv_class_eq [OF \equiv A r\]) qed subsection \Defining binary operations upon equivalence classes\ text \A congruence-preserving function of two arguments.\ definition congruent2 :: "('a \ 'a) set \ ('b \ 'b) set \ ('a \ 'b \ 'c) \ bool" where "congruent2 r1 r2 f \ (\(y1, z1) \ r1. \(y2, z2) \ r2. f y1 y2 = f z1 z2)" lemma congruent2I': assumes "\y1 z1 y2 z2. (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" shows "congruent2 r1 r2 f" using assms by (auto simp add: congruent2_def) lemma congruent2D: "congruent2 r1 r2 f \ (y1, z1) \ r1 \ (y2, z2) \ r2 \ f y1 y2 = f z1 z2" by (auto simp add: congruent2_def) text \Abbreviation for the common case where the relations are identical.\ abbreviation RESPECTS2:: "('a \ 'a \ 'b) \ ('a \ 'a) set \ bool" (infixr "respects2" 80) where "f respects2 r \ congruent2 r r f" lemma congruent2_implies_congruent: "equiv A r1 \ congruent2 r1 r2 f \ a \ A \ congruent r2 (f a)" unfolding congruent_def congruent2_def equiv_def refl_on_def by blast lemma congruent2_implies_congruent_UN: assumes "equiv A1 r1" "equiv A2 r2" "congruent2 r1 r2 f" "a \ A2" shows "congruent r1 (\x1. \x2 \ r2``{a}. f x1 x2)" unfolding congruent_def proof clarify fix c d assume cd: "(c,d) \ r1" then have "c \ A1" "d \ A1" using \equiv A1 r1\ by (auto elim!: equiv_type [THEN subsetD, THEN SigmaE2]) with assms show "\ (f c ` r2 `` {a}) = \ (f d ` r2 `` {a})" proof (simp add: UN_equiv_class congruent2_implies_congruent) show "f c a = f d a" using assms cd unfolding congruent2_def equiv_def refl_on_def by blast qed qed lemma UN_equiv_class2: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ a1 \ A1 \ a2 \ A2 \ (\x1 \ r1``{a1}. \x2 \ r2``{a2}. f x1 x2) = f a1 a2" by (simp add: UN_equiv_class congruent2_implies_congruent congruent2_implies_congruent_UN) lemma UN_equiv_class_type2: "equiv A1 r1 \ equiv A2 r2 \ congruent2 r1 r2 f \ X1 \ A1//r1 \ X2 \ A2//r2 \ (\x1 x2. x1 \ A1 \ x2 \ A2 \ f x1 x2 \ B) \ (\x1 \ X1. \x2 \ X2. f x1 x2) \ B" unfolding quotient_def by (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN congruent2_implies_congruent quotientI) lemma UN_UN_split_split_eq: "(\(x1, x2) \ X. \(y1, y2) \ Y. A x1 x2 y1 y2) = (\x \ X. \y \ Y. (\(x1, x2). (\(y1, y2). A x1 x2 y1 y2) y) x)" \ \Allows a natural expression of binary operators,\ \ \without explicit calls to \split\\ by auto lemma congruent2I: "equiv A1 r1 \ equiv A2 r2 \ (\y z w. w \ A2 \ (y,z) \ r1 \ f y w = f z w) \ (\y z w. w \ A1 \ (y,z) \ r2 \ f w y = f w z) \ congruent2 r1 r2 f" \ \Suggested by John Harrison -- the two subproofs may be\ \ \\<^emph>\much\ simpler than the direct proof.\ unfolding congruent2_def equiv_def refl_on_def by (blast intro: trans) lemma congruent2_commuteI: assumes equivA: "equiv A r" and commute: "\y z. y \ A \ z \ A \ f y z = f z y" and congt: "\y z w. w \ A \ (y,z) \ r \ f w y = f w z" shows "f respects2 r" proof (rule congruent2I [OF equivA equivA]) note eqv = equivA [THEN equiv_type, THEN subsetD, THEN SigmaE2] show "\y z w. \w \ A; (y, z) \ r\ \ f y w = f z w" by (iprover intro: commute [THEN trans] sym congt elim: eqv) show "\y z w. \w \ A; (y, z) \ r\ \ f w y = f w z" by (iprover intro: congt elim: eqv) qed subsection \Quotients and finiteness\ text \Suggested by Florian Kammüller\ lemma finite_quotient: assumes "finite A" "r \ A \ A" shows "finite (A//r)" \ \recall @{thm equiv_type}\ proof - have "A//r \ Pow A" using assms unfolding quotient_def by blast moreover have "finite (Pow A)" using assms by simp ultimately show ?thesis by (iprover intro: finite_subset) qed lemma finite_equiv_class: "finite A \ r \ A \ A \ X \ A//r \ finite X" unfolding quotient_def by (erule rev_finite_subset) blast lemma equiv_imp_dvd_card: assumes "finite A" "equiv A r" "\X. X \ A//r \ k dvd card X" shows "k dvd card A" proof (rule Union_quotient [THEN subst]) show "k dvd card (\ (A // r))" apply (rule dvd_partition) using assms by (auto simp: Union_quotient dest: quotient_disj) qed (use assms in blast) lemma card_quotient_disjoint: assumes "finite A" "inj_on (\x. {x} // r) A" shows "card (A//r) = card A" proof - have "\i\A. \j\A. i \ j \ r `` {j} \ r `` {i}" using assms by (fastforce simp add: quotient_def inj_on_def) with assms show ?thesis by (simp add: quotient_def card_UN_disjoint) qed text \By Jakub Kądziołka:\ lemma sum_fun_comp: assumes "finite S" "finite R" "g ` S \ R" shows "(\x \ S. f (g x)) = (\y \ R. of_nat (card {x \ S. g x = y}) * f y)" proof - let ?r = "relation_of (\p q. g p = g q) S" have eqv: "equiv S ?r" unfolding relation_of_def by (auto intro: comp_equivI) have finite: "C \ S//?r \ finite C" for C - by (fact finite_equiv_class[OF `finite S` equiv_type[OF `equiv S ?r`]]) + by (fact finite_equiv_class[OF \finite S\ equiv_type[OF \equiv S ?r\]]) have disjoint: "A \ S//?r \ B \ S//?r \ A \ B \ A \ B = {}" for A B using eqv quotient_disj by blast let ?cls = "\y. {x \ S. y = g x}" have quot_as_img: "S//?r = ?cls ` g ` S" by (auto simp add: relation_of_def quotient_def) have cls_inj: "inj_on ?cls (g ` S)" by (auto intro: inj_onI) have rest_0: "(\y \ R - g ` S. of_nat (card (?cls y)) * f y) = 0" proof - have "of_nat (card (?cls y)) * f y = 0" if asm: "y \ R - g ` S" for y proof - from asm have *: "?cls y = {}" by auto show ?thesis unfolding * by simp qed thus ?thesis by simp qed have "(\x \ S. f (g x)) = (\C \ S//?r. \x \ C. f (g x))" using eqv finite disjoint by (simp flip: sum.Union_disjoint[simplified] add: Union_quotient) also have "... = (\y \ g ` S. \x \ ?cls y. f (g x))" unfolding quot_as_img by (simp add: sum.reindex[OF cls_inj]) also have "... = (\y \ g ` S. \x \ ?cls y. f y)" by auto also have "... = (\y \ g ` S. of_nat (card (?cls y)) * f y)" by (simp flip: sum_constant) also have "... = (\y \ R. of_nat (card (?cls y)) * f y)" using rest_0 by (simp add: sum.subset_diff[OF \g ` S \ R\ \finite R\]) finally show ?thesis by (simp add: eq_commute) qed subsection \Projection\ definition proj :: "('b \ 'a) set \ 'b \ 'a set" where "proj r x = r `` {x}" lemma proj_preserves: "x \ A \ proj r x \ A//r" unfolding proj_def by (rule quotientI) lemma proj_in_iff: assumes "equiv A r" shows "proj r x \ A//r \ x \ A" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: proj_preserves) next assume ?lhs then show ?rhs unfolding proj_def quotient_def proof clarsimp fix y assume y: "y \ A" and "r `` {x} = r `` {y}" moreover have "y \ r `` {y}" using assms y unfolding equiv_def refl_on_def by blast ultimately have "(x, y) \ r" by blast then show "x \ A" using assms unfolding equiv_def refl_on_def by blast qed qed lemma proj_iff: "equiv A r \ {x, y} \ A \ proj r x = proj r y \ (x, y) \ r" by (simp add: proj_def eq_equiv_class_iff) (* lemma in_proj: "\equiv A r; x \ A\ \ x \ proj r x" unfolding proj_def equiv_def refl_on_def by blast *) lemma proj_image: "proj r ` A = A//r" unfolding proj_def[abs_def] quotient_def by blast lemma in_quotient_imp_non_empty: "equiv A r \ X \ A//r \ X \ {}" unfolding quotient_def using equiv_class_self by fast lemma in_quotient_imp_in_rel: "equiv A r \ X \ A//r \ {x, y} \ X \ (x, y) \ r" using quotient_eq_iff[THEN iffD1] by fastforce lemma in_quotient_imp_closed: "equiv A r \ X \ A//r \ x \ X \ (x, y) \ r \ y \ X" unfolding quotient_def equiv_def trans_def by blast lemma in_quotient_imp_subset: "equiv A r \ X \ A//r \ X \ A" using in_quotient_imp_in_rel equiv_type by fastforce subsection \Equivalence relations -- predicate version\ text \Partial equivalences.\ definition part_equivp :: "('a \ 'a \ bool) \ bool" where "part_equivp R \ (\x. R x x) \ (\x y. R x y \ R x x \ R y y \ R x = R y)" \ \John-Harrison-style characterization\ lemma part_equivpI: "\x. R x x \ symp R \ transp R \ part_equivp R" by (auto simp add: part_equivp_def) (auto elim: sympE transpE) lemma part_equivpE: assumes "part_equivp R" obtains x where "R x x" and "symp R" and "transp R" proof - from assms have 1: "\x. R x x" and 2: "\x y. R x y \ R x x \ R y y \ R x = R y" unfolding part_equivp_def by blast+ from 1 obtain x where "R x x" .. moreover have "symp R" proof (rule sympI) fix x y assume "R x y" with 2 [of x y] show "R y x" by auto qed moreover have "transp R" proof (rule transpI) fix x y z assume "R x y" and "R y z" with 2 [of x y] 2 [of y z] show "R x z" by auto qed ultimately show thesis by (rule that) qed lemma part_equivp_refl_symp_transp: "part_equivp R \ (\x. R x x) \ symp R \ transp R" by (auto intro: part_equivpI elim: part_equivpE) lemma part_equivp_symp: "part_equivp R \ R x y \ R y x" by (erule part_equivpE, erule sympE) lemma part_equivp_transp: "part_equivp R \ R x y \ R y z \ R x z" by (erule part_equivpE, erule transpE) lemma part_equivp_typedef: "part_equivp R \ \d. d \ {c. \x. R x x \ c = Collect (R x)}" by (auto elim: part_equivpE) text \Total equivalences.\ definition equivp :: "('a \ 'a \ bool) \ bool" where "equivp R \ (\x y. R x y = (R x = R y))" \ \John-Harrison-style characterization\ lemma equivpI: "reflp R \ symp R \ transp R \ equivp R" by (auto elim: reflpE sympE transpE simp add: equivp_def) lemma equivpE: assumes "equivp R" obtains "reflp R" and "symp R" and "transp R" using assms by (auto intro!: that reflpI sympI transpI simp add: equivp_def) lemma equivp_implies_part_equivp: "equivp R \ part_equivp R" by (auto intro: part_equivpI elim: equivpE reflpE) lemma equivp_equiv: "equiv UNIV A \ equivp (\x y. (x, y) \ A)" by (auto intro!: equivI equivpI [to_set] elim!: equivE equivpE [to_set]) lemma equivp_reflp_symp_transp: "equivp R \ reflp R \ symp R \ transp R" by (auto intro: equivpI elim: equivpE) lemma identity_equivp: "equivp (=)" by (auto intro: equivpI reflpI sympI transpI) lemma equivp_reflp: "equivp R \ R x x" by (erule equivpE, erule reflpE) lemma equivp_symp: "equivp R \ R x y \ R y x" by (erule equivpE, erule sympE) lemma equivp_transp: "equivp R \ R x y \ R y z \ R x z" by (erule equivpE, erule transpE) lemma equivp_rtranclp: "symp r \ equivp r\<^sup>*\<^sup>*" by(intro equivpI reflpI sympI transpI)(auto dest: sympD[OF symp_rtranclp]) lemmas equivp_rtranclp_symclp [simp] = equivp_rtranclp[OF symp_symclp] lemma equivp_vimage2p: "equivp R \ equivp (vimage2p f f R)" by(auto simp add: equivp_def vimage2p_def dest: fun_cong) lemma equivp_imp_transp: "equivp R \ transp R" by(simp add: equivp_reflp_symp_transp) subsection \Equivalence closure\ definition equivclp :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" where "equivclp r = (symclp r)\<^sup>*\<^sup>*" lemma transp_equivclp [simp]: "transp (equivclp r)" by(simp add: equivclp_def) lemma reflp_equivclp [simp]: "reflp (equivclp r)" by(simp add: equivclp_def) lemma symp_equivclp [simp]: "symp (equivclp r)" by(simp add: equivclp_def) lemma equivp_evquivclp [simp]: "equivp (equivclp r)" by(simp add: equivpI) lemma tranclp_equivclp [simp]: "(equivclp r)\<^sup>+\<^sup>+ = equivclp r" by(simp add: equivclp_def) lemma rtranclp_equivclp [simp]: "(equivclp r)\<^sup>*\<^sup>* = equivclp r" by(simp add: equivclp_def) lemma symclp_equivclp [simp]: "symclp (equivclp r) = equivclp r" by(simp add: equivclp_def symp_symclp_eq) lemma equivclp_symclp [simp]: "equivclp (symclp r) = equivclp r" by(simp add: equivclp_def) lemma equivclp_conversep [simp]: "equivclp (conversep r) = equivclp r" by(simp add: equivclp_def) lemma equivclp_sym [sym]: "equivclp r x y \ equivclp r y x" by(rule sympD[OF symp_equivclp]) lemma equivclp_OO_equivclp_le_equivclp: "equivclp r OO equivclp r \ equivclp r" by(rule transp_relcompp_less_eq transp_equivclp)+ lemma rtranlcp_le_equivclp: "r\<^sup>*\<^sup>* \ equivclp r" unfolding equivclp_def by(rule rtranclp_mono)(simp add: symclp_pointfree) lemma rtranclp_conversep_le_equivclp: "r\\\<^sup>*\<^sup>* \ equivclp r" unfolding equivclp_def by(rule rtranclp_mono)(simp add: symclp_pointfree) lemma symclp_rtranclp_le_equivclp: "symclp r\<^sup>*\<^sup>* \ equivclp r" unfolding symclp_pointfree by(rule le_supI)(simp_all add: rtranclp_conversep[symmetric] rtranlcp_le_equivclp rtranclp_conversep_le_equivclp) lemma r_OO_conversep_into_equivclp: "r\<^sup>*\<^sup>* OO r\\\<^sup>*\<^sup>* \ equivclp r" by(blast intro: order_trans[OF _ equivclp_OO_equivclp_le_equivclp] relcompp_mono rtranlcp_le_equivclp rtranclp_conversep_le_equivclp del: predicate2I) lemma equivclp_induct [consumes 1, case_names base step, induct pred: equivclp]: assumes a: "equivclp r a b" and cases: "P a" "\y z. equivclp r a y \ r y z \ r z y \ P y \ P z" shows "P b" using a unfolding equivclp_def by(induction rule: rtranclp_induct; fold equivclp_def; blast intro: cases elim: symclpE) lemma converse_equivclp_induct [consumes 1, case_names base step]: assumes major: "equivclp r a b" and cases: "P b" "\y z. r y z \ r z y \ equivclp r z b \ P z \ P y" shows "P a" using major unfolding equivclp_def by(induction rule: converse_rtranclp_induct; fold equivclp_def; blast intro: cases elim: symclpE) lemma equivclp_refl [simp]: "equivclp r x x" by(rule reflpD[OF reflp_equivclp]) lemma r_into_equivclp [intro]: "r x y \ equivclp r x y" unfolding equivclp_def by(blast intro: symclpI) lemma converse_r_into_equivclp [intro]: "r y x \ equivclp r x y" unfolding equivclp_def by(blast intro: symclpI) lemma rtranclp_into_equivclp: "r\<^sup>*\<^sup>* x y \ equivclp r x y" using rtranlcp_le_equivclp[of r] by blast lemma converse_rtranclp_into_equivclp: "r\<^sup>*\<^sup>* y x \ equivclp r x y" by(blast intro: equivclp_sym rtranclp_into_equivclp) lemma equivclp_into_equivclp: "\ equivclp r a b; r b c \ r c b \ \ equivclp r a c" unfolding equivclp_def by(erule rtranclp.rtrancl_into_rtrancl)(auto intro: symclpI) lemma equivclp_trans [trans]: "\ equivclp r a b; equivclp r b c \ \ equivclp r a c" using equivclp_OO_equivclp_le_equivclp[of r] by blast hide_const (open) proj end diff --git a/src/HOL/Set_Interval.thy b/src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy +++ b/src/HOL/Set_Interval.thy @@ -1,2564 +1,2564 @@ (* Title: HOL/Set_Interval.thy Author: Tobias Nipkow, Clemens Ballarin, Jeremy Avigad lessThan, greaterThan, atLeast, atMost and two-sided intervals Modern convention: Ixy stands for an interval where x and y describe the lower and upper bound and x,y : {c,o,i} where c = closed, o = open, i = infinite. Examples: Ico = {_ ..< _} and Ici = {_ ..} *) section \Set intervals\ theory Set_Interval imports Divides begin (* Belongs in Finite_Set but 2 is not available there *) lemma card_2_iff: "card S = 2 \ (\x y. S = {x,y} \ x \ y)" by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_2_iff': "card S = 2 \ (\x\S. \y\S. x \ y \ (\z\S. z = x \ z = y))" by (auto simp: card_Suc_eq numeral_eq_Suc) context ord begin definition lessThan :: "'a => 'a set" ("(1{..<_})") where "{.. 'a set" ("(1{.._})") where "{..u} == {x. x \ u}" definition greaterThan :: "'a => 'a set" ("(1{_<..})") where "{l<..} == {x. l 'a set" ("(1{_..})") where "{l..} == {x. l\x}" definition greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where "{l<.. 'a => 'a set" ("(1{_..<_})") where "{l.. 'a => 'a set" ("(1{_<.._})") where "{l<..u} == {l<..} Int {..u}" definition atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where "{l..u} == {l..} Int {..u}" end text\A note of warning when using \<^term>\{.. on type \<^typ>\nat\: it is equivalent to \<^term>\{0::nat.. but some lemmas involving \<^term>\{m.. may not exist in \<^term>\{..-form as well.\ syntax (ASCII) "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3UN _<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3INT _<=_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3INT _<_./ _)" [0, 0, 10] 10) syntax (latex output) "_UNION_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_UNION_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) "_INTER_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_INTER_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) syntax "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) translations "\i\n. A" \ "\i\{..n}. A" "\i "\i\{..i\n. A" \ "\i\{..n}. A" "\i "\i\{..Various equivalences\ lemma (in ord) lessThan_iff [iff]: "(i \ lessThan k) = (i greaterThan k) = (k atLeast k) = (k<=i)" by (simp add: atLeast_def) lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k" by (auto simp add: lessThan_def atLeast_def) lemma (in ord) atMost_iff [iff]: "(i \ atMost k) = (i<=k)" by (simp add: atMost_def) lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}" by (blast intro: order_antisym) lemma (in linorder) lessThan_Int_lessThan: "{ a <..} \ { b <..} = { max a b <..}" by auto lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \ {..< b} = {..< min a b}" by auto subsection \Logical Equivalences for Set Inclusion and Equality\ lemma atLeast_empty_triv [simp]: "{{}..} = UNIV" by auto lemma atMost_UNIV_triv [simp]: "{..UNIV} = UNIV" by auto lemma atLeast_subset_iff [iff]: "(atLeast x \ atLeast y) = (y \ (x::'a::preorder))" by (blast intro: order_trans) lemma atLeast_eq_iff [iff]: "(atLeast x = atLeast y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma greaterThan_subset_iff [iff]: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" unfolding greaterThan_def by (auto simp: linorder_not_less [symmetric]) lemma greaterThan_eq_iff [iff]: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::preorder))" by (blast intro: order_trans) lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma lessThan_subset_iff [iff]: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" unfolding lessThan_def by (auto simp: linorder_not_less [symmetric]) lemma lessThan_eq_iff [iff]: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma lessThan_strict_subset_iff: fixes m n :: "'a::linorder" shows "{.. m < n" by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq) lemma (in linorder) Ici_subset_Ioi_iff: "{a ..} \ {b <..} \ b < a" by auto lemma (in linorder) Iic_subset_Iio_iff: "{.. a} \ {..< b} \ a < b" by auto lemma (in preorder) Ioi_le_Ico: "{a <..} \ {a ..}" by (auto intro: less_imp_le) subsection \Two-sided intervals\ context ord begin lemma greaterThanLessThan_iff [simp]: "(i \ {l<.. i < u)" by (simp add: greaterThanLessThan_def) lemma atLeastLessThan_iff [simp]: "(i \ {l.. i \ i < u)" by (simp add: atLeastLessThan_def) lemma greaterThanAtMost_iff [simp]: "(i \ {l<..u}) = (l < i \ i \ u)" by (simp add: greaterThanAtMost_def) lemma atLeastAtMost_iff [simp]: "(i \ {l..u}) = (l \ i \ i \ u)" by (simp add: atLeastAtMost_def) text \The above four lemmas could be declared as iffs. Unfortunately this breaks many proofs. Since it only helps blast, it is better to leave them alone.\ lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \ {..< b }" by auto lemma (in order) atLeastLessThan_eq_atLeastAtMost_diff: "{a..Emptyness, singletons, subset\ context preorder begin lemma atLeastatMost_empty_iff[simp]: "{a..b} = {} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastatMost_empty_iff2[simp]: "{} = {a..b} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastLessThan_empty_iff[simp]: "{a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma atLeastLessThan_empty_iff2[simp]: "{} = {a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \ \ k < l" by auto (blast intro: less_le_trans) lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \ \ k < l" by auto (blast intro: less_le_trans) lemma atLeastatMost_subset_iff[simp]: "{a..b} \ {c..d} \ (\ a \ b) \ c \ a \ b \ d" unfolding atLeastAtMost_def atLeast_def atMost_def by (blast intro: order_trans) lemma atLeastatMost_psubset_iff: "{a..b} < {c..d} \ ((\ a \ b) \ c \ a \ b \ d \ (c < a \ b < d)) \ c \ d" by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans) lemma atLeastAtMost_subseteq_atLeastLessThan_iff: "{a..b} \ {c ..< d} \ (a \ b \ c \ a \ b < d)" by auto (blast intro: local.order_trans local.le_less_trans elim: )+ lemma Icc_subset_Ici_iff[simp]: "{l..h} \ {l'..} = (\ l\h \ l\l')" by(auto simp: subset_eq intro: order_trans) lemma Icc_subset_Iic_iff[simp]: "{l..h} \ {..h'} = (\ l\h \ h\h')" by(auto simp: subset_eq intro: order_trans) lemma not_Ici_eq_empty[simp]: "{l..} \ {}" by(auto simp: set_eq_iff) lemma not_Iic_eq_empty[simp]: "{..h} \ {}" by(auto simp: set_eq_iff) lemmas not_empty_eq_Ici_eq_empty[simp] = not_Ici_eq_empty[symmetric] lemmas not_empty_eq_Iic_eq_empty[simp] = not_Iic_eq_empty[symmetric] end context order begin lemma atLeastatMost_empty[simp]: "b < a \ {a..b} = {}" by(auto simp: atLeastAtMost_def atLeast_def atMost_def) lemma atLeastLessThan_empty[simp]: "b \ a \ {a.. k ==> {k<..l} = {}" by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) lemma greaterThanLessThan_empty[simp]:"l \ k ==> {k<.. {a .. b} = {a}" by simp lemma Icc_eq_Icc[simp]: "{l..h} = {l'..h'} = (l=l' \ h=h' \ \ l\h \ \ l'\h')" by (simp add: order_class.order.eq_iff) (auto intro: order_trans) lemma atLeastAtMost_singleton_iff[simp]: "{a .. b} = {c} \ a = b \ b = c" proof assume "{a..b} = {c}" hence *: "\ (\ a \ b)" unfolding atLeastatMost_empty_iff[symmetric] by simp with \{a..b} = {c}\ have "c \ a \ b \ c" by auto with * show "a = b \ b = c" by auto qed simp end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_le_Icc[simp]: "\ UNIV \ {l..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_UNIV_le_Iic[simp]: "\ UNIV \ {..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_Ici_le_Icc[simp]: "\ {l..} \ {l'..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Ici_le_Iic[simp]: "\ {l..} \ {..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_bot begin lemma not_UNIV_le_Ici[simp]: "\ UNIV \ {l..}" using lt_ex[of l] by(auto simp: subset_eq less_le_not_le) lemma not_Iic_le_Icc[simp]: "\ {..h} \ {l'..h'}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Iic_le_Ici[simp]: "\ {..h} \ {l'..}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_eq_Icc[simp]: "\ UNIV = {l'..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Icc_eq_UNIV[simp] = not_UNIV_eq_Icc[symmetric] lemma not_UNIV_eq_Iic[simp]: "\ UNIV = {..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Iic_eq_UNIV[simp] = not_UNIV_eq_Iic[symmetric] lemma not_Icc_eq_Ici[simp]: "\ {l..h} = {l'..}" unfolding atLeastAtMost_def using not_Ici_le_Iic[of l'] by blast lemmas not_Ici_eq_Icc[simp] = not_Icc_eq_Ici[symmetric] (* also holds for no_bot but no_top should suffice *) lemma not_Iic_eq_Ici[simp]: "\ {..h} = {l'..}" using not_Ici_le_Iic[of l' h] by blast lemmas not_Ici_eq_Iic[simp] = not_Iic_eq_Ici[symmetric] end context no_bot begin lemma not_UNIV_eq_Ici[simp]: "\ UNIV = {l'..}" using lt_ex[of l'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Ici_eq_UNIV[simp] = not_UNIV_eq_Ici[symmetric] lemma not_Icc_eq_Iic[simp]: "\ {l..h} = {..h'}" unfolding atLeastAtMost_def using not_Iic_le_Ici[of h'] by blast lemmas not_Iic_eq_Icc[simp] = not_Icc_eq_Iic[symmetric] end context dense_linorder begin lemma greaterThanLessThan_empty_iff[simp]: "{ a <..< b } = {} \ b \ a" using dense[of a b] by (cases "a < b") auto lemma greaterThanLessThan_empty_iff2[simp]: "{} = { a <..< b } \ b \ a" using dense[of a b] by (cases "a < b") auto lemma atLeastLessThan_subseteq_atLeastAtMost_iff: "{a ..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastAtMost_iff: "{a <.. b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastAtMost_iff: "{a <..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanLessThan: "{a <..< b} \ {c <..< d} \ (a < b \ a \ c \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastLessThan_iff: "{a <.. b} \ { c ..< d } \ (a < b \ c \ a \ b < d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastLessThan_iff: "{a <..< b} \ { c ..< d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff: "{a <..< b} \ { c <.. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) end context no_top begin lemma greaterThan_non_empty[simp]: "{x <..} \ {}" using gt_ex[of x] by auto end context no_bot begin lemma lessThan_non_empty[simp]: "{..< x} \ {}" using lt_ex[of x] by auto end lemma (in linorder) atLeastLessThan_subset_iff: "{a.. {c.. b \ a \ c\a \ b\d" apply (auto simp:subset_eq Ball_def not_le) apply(frule_tac x=a in spec) apply(erule_tac x=d in allE) apply (auto simp: ) done lemma atLeastLessThan_inj: fixes a b c d :: "'a::linorder" assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d" shows "a = c" "b = d" using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le antisym_conv2 subset_refl)+ lemma atLeastLessThan_eq_iff: fixes a b c d :: "'a::linorder" assumes "a < b" "c < d" shows "{a ..< b} = {c ..< d} \ a = c \ b = d" using atLeastLessThan_inj assms by auto lemma (in linorder) Ioc_inj: \{a <.. b} = {c <.. d} \ (b \ a \ d \ c) \ a = c \ b = d\ (is \?P \ ?Q\) proof assume ?Q then show ?P by auto next assume ?P then have \a < x \ x \ b \ c < x \ x \ d\ for x by (simp add: set_eq_iff) from this [of a] this [of b] this [of c] this [of d] show ?Q by auto qed lemma (in order) Iio_Int_singleton: "{.. {x} = (if x < k then {x} else {})" by auto lemma (in linorder) Ioc_subset_iff: "{a<..b} \ {c<..d} \ (b \ a \ c \ a \ b \ d)" by (auto simp: subset_eq Ball_def) (metis less_le not_less) lemma (in order_bot) atLeast_eq_UNIV_iff: "{x..} = UNIV \ x = bot" by (auto simp: set_eq_iff intro: le_bot) lemma (in order_top) atMost_eq_UNIV_iff: "{..x} = UNIV \ x = top" by (auto simp: set_eq_iff intro: top_le) lemma (in bounded_lattice) atLeastAtMost_eq_UNIV_iff: "{x..y} = UNIV \ (x = bot \ y = top)" by (auto simp: set_eq_iff intro: top_le le_bot) lemma Iio_eq_empty_iff: "{..< n::'a::{linorder, order_bot}} = {} \ n = bot" by (auto simp: set_eq_iff not_less le_bot) lemma lessThan_empty_iff: "{..< n::nat} = {} \ n = 0" by (simp add: Iio_eq_empty_iff bot_nat_def) lemma mono_image_least: assumes f_mono: "mono f" and f_img: "f ` {m ..< n} = {m' ..< n'}" "m < n" shows "f m = m'" proof - from f_img have "{m' ..< n'} \ {}" by (metis atLeastLessThan_empty_iff image_is_empty) with f_img have "m' \ f ` {m ..< n}" by auto then obtain k where "f k = m'" "m \ k" by auto moreover have "m' \ f m" using f_img by auto ultimately show "f m = m'" using f_mono by (auto elim: monoE[where x=m and y=k]) qed subsection \Infinite intervals\ context dense_linorder begin lemma infinite_Ioo: assumes "a < b" shows "\ finite {a<.. {}" using \a < b\ by auto ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b" using Max_in[of "{a <..< b}"] by auto then obtain x where "Max {a <..< b} < x" "x < b" using dense[of "Max {a<.. {a <..< b}" using \a < Max {a <..< b}\ by auto then have "x \ Max {a <..< b}" using fin by auto with \Max {a <..< b} < x\ show False by auto qed lemma infinite_Icc: "a < b \ \ finite {a .. b}" using greaterThanLessThan_subseteq_atLeastAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ico: "a < b \ \ finite {a ..< b}" using greaterThanLessThan_subseteq_atLeastLessThan_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioc: "a < b \ \ finite {a <.. b}" using greaterThanLessThan_subseteq_greaterThanAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioo_iff [simp]: "infinite {a<.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioo) lemma infinite_Icc_iff [simp]: "infinite {a .. b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Icc) lemma infinite_Ico_iff [simp]: "infinite {a.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ico) lemma infinite_Ioc_iff [simp]: "infinite {a<..b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioc) end lemma infinite_Iio: "\ finite {..< a :: 'a :: {no_bot, linorder}}" proof assume "finite {..< a}" then have *: "\x. x < a \ Min {..< a} \ x" by auto obtain x where "x < a" using lt_ex by auto obtain y where "y < Min {..< a}" using lt_ex by auto also have "Min {..< a} \ x" using \x < a\ by fact also note \x < a\ finally have "Min {..< a} \ y" by fact with \y < Min {..< a}\ show False by auto qed lemma infinite_Iic: "\ finite {.. a :: 'a :: {no_bot, linorder}}" using infinite_Iio[of a] finite_subset[of "{..< a}" "{.. a}"] by (auto simp: subset_eq less_imp_le) lemma infinite_Ioi: "\ finite {a :: 'a :: {no_top, linorder} <..}" proof assume "finite {a <..}" then have *: "\x. a < x \ x \ Max {a <..}" by auto obtain y where "Max {a <..} < y" using gt_ex by auto obtain x where x: "a < x" using gt_ex by auto also from x have "x \ Max {a <..}" by fact also note \Max {a <..} < y\ finally have "y \ Max { a <..}" by fact with \Max {a <..} < y\ show False by auto qed lemma infinite_Ici: "\ finite {a :: 'a :: {no_top, linorder} ..}" using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"] by (auto simp: subset_eq less_imp_le) subsubsection \Intersection\ context linorder begin lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}" by auto lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}" by auto lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}" by auto lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}" by auto lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}" by auto lemma Int_atLeastLessThan[simp]: "{a.. {..b} = {.. min a b}" by (auto simp: min_def) lemma Ioc_disjoint: "{a<..b} \ {c<..d} = {} \ b \ a \ d \ c \ b \ c \ d \ a" by auto end context complete_lattice begin lemma shows Sup_atLeast[simp]: "Sup {x ..} = top" and Sup_greaterThanAtLeast[simp]: "x < top \ Sup {x <..} = top" and Sup_atMost[simp]: "Sup {.. y} = y" and Sup_atLeastAtMost[simp]: "x \ y \ Sup { x .. y} = y" and Sup_greaterThanAtMost[simp]: "x < y \ Sup { x <.. y} = y" by (auto intro!: Sup_eqI) lemma shows Inf_atMost[simp]: "Inf {.. x} = bot" and Inf_atMostLessThan[simp]: "top < x \ Inf {..< x} = bot" and Inf_atLeast[simp]: "Inf {x ..} = x" and Inf_atLeastAtMost[simp]: "x \ y \ Inf { x .. y} = x" and Inf_atLeastLessThan[simp]: "x < y \ Inf { x ..< y} = x" by (auto intro!: Inf_eqI) end lemma fixes x y :: "'a :: {complete_lattice, dense_linorder}" shows Sup_lessThan[simp]: "Sup {..< y} = y" and Sup_atLeastLessThan[simp]: "x < y \ Sup { x ..< y} = y" and Sup_greaterThanLessThan[simp]: "x < y \ Sup { x <..< y} = y" and Inf_greaterThan[simp]: "Inf {x <..} = x" and Inf_greaterThanAtMost[simp]: "x < y \ Inf { x <.. y} = x" and Inf_greaterThanLessThan[simp]: "x < y \ Inf { x <..< y} = x" by (auto intro!: Inf_eqI Sup_eqI intro: dense_le dense_le_bounded dense_ge dense_ge_bounded) subsection \Intervals of natural numbers\ subsubsection \The Constant \<^term>\lessThan\\ lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" by (simp add: lessThan_def) lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" by (simp add: lessThan_def less_Suc_eq, blast) text \The following proof is convenient in induction proofs where new elements get indices at the beginning. So it is used to transform \<^term>\{.. to \<^term>\0::nat\ and \<^term>\{..< n}\.\ lemma zero_notin_Suc_image [simp]: "0 \ Suc ` A" by auto lemma lessThan_Suc_eq_insert_0: "{..m::nat. lessThan m) = UNIV" by blast subsubsection \The Constant \<^term>\greaterThan\\ lemma greaterThan_0: "greaterThan 0 = range Suc" unfolding greaterThan_def by (blast dest: gr0_conv_Suc [THEN iffD1]) lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" unfolding greaterThan_def by (auto elim: linorder_neqE) lemma INT_greaterThan_UNIV: "(\m::nat. greaterThan m) = {}" by blast subsubsection \The Constant \<^term>\atLeast\\ lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" by (unfold atLeast_def UNIV_def, simp) lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" unfolding atLeast_def by (auto simp: order_le_less Suc_le_eq) lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) lemma UN_atLeast_UNIV: "(\m::nat. atLeast m) = UNIV" by blast subsubsection \The Constant \<^term>\atMost\\ lemma atMost_0 [simp]: "atMost (0::nat) = {0}" by (simp add: atMost_def) lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" unfolding atMost_def by (auto simp add: less_Suc_eq order_le_less) lemma UN_atMost_UNIV: "(\m::nat. atMost m) = UNIV" by blast subsubsection \The Constant \<^term>\atLeastLessThan\\ text\The orientation of the following 2 rules is tricky. The lhs is defined in terms of the rhs. Hence the chosen orientation makes sense in this theory --- the reverse orientation complicates proofs (eg nontermination). But outside, when the definition of the lhs is rarely used, the opposite orientation seems preferable because it reduces a specific concept to a more general one.\ lemma atLeast0LessThan [code_abbrev]: "{0::nat..The Constant \<^term>\atLeastAtMost\\ lemma Icc_eq_insert_lb_nat: "m \ n \ {m..n} = insert m {Suc m..n}" by auto lemma atLeast0_atMost_Suc: "{0..Suc n} = insert (Suc n) {0..n}" by (simp add: atLeast0AtMost atMost_Suc) lemma atLeast0_atMost_Suc_eq_insert_0: "{0..Suc n} = insert 0 (Suc ` {0..n})" by (simp add: atLeast0AtMost atMost_Suc_eq_insert_0) subsubsection \Intervals of nats with \<^term>\Suc\\ text\Not a simprule because the RHS is too messy.\ lemma atLeastLessThanSuc: "{m.. n then insert n {m.. Suc n \ {m..Suc n} = insert (Suc n) {m..n}" by auto lemma atLeastAtMost_insertL: "m \ n \ insert m {Suc m..n} = {m ..n}" by auto text \The analogous result is useful on \<^typ>\int\:\ (* here, because we don't have an own int section *) lemma atLeastAtMostPlus1_int_conv: "m \ 1+n \ {m..1+n} = insert (1+n) {m..n::int}" by (auto intro: set_eqI) lemma atLeastLessThan_add_Un: "i \ j \ {i.. {j..Intervals and numerals\ lemma lessThan_nat_numeral: \ \Evaluation for specific numerals\ "lessThan (numeral k :: nat) = insert (pred_numeral k) (lessThan (pred_numeral k))" by (simp add: numeral_eq_Suc lessThan_Suc) lemma atMost_nat_numeral: \ \Evaluation for specific numerals\ "atMost (numeral k :: nat) = insert (numeral k) (atMost (pred_numeral k))" by (simp add: numeral_eq_Suc atMost_Suc) lemma atLeastLessThan_nat_numeral: \ \Evaluation for specific numerals\ "atLeastLessThan m (numeral k :: nat) = (if m \ (pred_numeral k) then insert (pred_numeral k) (atLeastLessThan m (pred_numeral k)) else {})" by (simp add: numeral_eq_Suc atLeastLessThanSuc) subsubsection \Image\ context linordered_semidom begin lemma image_add_atLeast[simp]: "plus k ` {i..} = {k + i..}" proof - have "n = k + (n - k)" if "i + k \ n" for n proof - have "n = (n - (k + i)) + (k + i)" using that by (metis add_commute le_add_diff_inverse) then show "n = k + (n - k)" by (metis local.add_diff_cancel_left' add_assoc add_commute) qed then show ?thesis by (fastforce simp: add_le_imp_le_diff add.commute) qed lemma image_add_atLeastAtMost [simp]: "plus k ` {i..j} = {i + k..j + k}" (is "?A = ?B") proof show "?A \ ?B" by (auto simp add: ac_simps) next show "?B \ ?A" proof fix n assume "n \ ?B" then have "i \ n - k" by (simp add: add_le_imp_le_diff) have "n = n - k + k" proof - from \n \ ?B\ have "n = n - (i + k) + (i + k)" by simp also have "\ = n - k - i + i + k" by (simp add: algebra_simps) also have "\ = n - k + k" using \i \ n - k\ by simp finally show ?thesis . qed moreover have "n - k \ {i..j}" using \n \ ?B\ by (auto simp: add_le_imp_le_diff add_le_add_imp_diff_le) ultimately show "n \ ?A" by (simp add: ac_simps) qed qed lemma image_add_atLeastAtMost' [simp]: "(\n. n + k) ` {i..j} = {i + k..j + k}" by (simp add: add.commute [of _ k]) lemma image_add_atLeastLessThan [simp]: "plus k ` {i..n. n + k) ` {i.. uminus ` {x<..}" by (rule imageI) (simp add: *) thus "y \ uminus ` {x<..}" by simp next fix y assume "y \ -x" have "- (-y) \ uminus ` {x..}" by (rule imageI) (insert \y \ -x\[THEN le_imp_neg_le], simp) thus "y \ uminus ` {x..}" by simp qed simp_all lemma fixes x :: 'a shows image_uminus_lessThan[simp]: "uminus ` {.. = {c - b<..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_greaterThanAtMost[simp]: fixes a b c::"'a::linordered_idom" shows "(-) c ` {a<..b} = {c - b.. = {c - b.. = {..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_AtMost[simp]: fixes b c::"'a::linordered_idom" shows "(-) c ` {..b} = {c - b..}" proof - have "(-) c ` {..b} = (+) c ` uminus ` {..b}" unfolding image_image by simp also have "\ = {c - b..}" by simp finally show ?thesis by simp qed lemma image_minus_const_atLeastAtMost' [simp]: "(\t. t-d)`{a..b} = {a-d..b-d}" for d::"'a::linordered_idom" by (metis (no_types, lifting) diff_conv_add_uminus image_add_atLeastAtMost' image_cong) context linordered_field begin lemma image_mult_atLeastAtMost [simp]: "((*) d ` {a..b}) = {d*a..d*b}" if "d>0" using that by (auto simp: field_simps mult_le_cancel_right intro: rev_image_eqI [where x="x/d" for x]) lemma image_divide_atLeastAtMost [simp]: "((\c. c / d) ` {a..b}) = {a/d..b/d}" if "d>0" proof - from that have "inverse d > 0" by simp with image_mult_atLeastAtMost [of "inverse d" a b] have "(*) (inverse d) ` {a..b} = {inverse d * a..inverse d * b}" by blast moreover have "(*) (inverse d) = (\c. c / d)" by (simp add: fun_eq_iff field_simps) ultimately show ?thesis by simp qed lemma image_mult_atLeastAtMost_if: "(*) c ` {x .. y} = (if c > 0 then {c * x .. c * y} else if x \ y then {c * y .. c * x} else {})" proof (cases "c = 0 \ x > y") case True then show ?thesis by auto next case False then have "x \ y" by auto from False consider "c < 0"| "c > 0" by (auto simp add: neq_iff) then show ?thesis proof cases case 1 have "(*) c ` {x..y} = {c * y..c * x}" proof (rule set_eqI) fix d from 1 have "inj (\z. z / c)" by (auto intro: injI) then have "d \ (*) c ` {x..y} \ d / c \ (\z. z div c) ` (*) c ` {x..y}" by (subst inj_image_mem_iff) simp_all also have "\ \ d / c \ {x..y}" using 1 by (simp add: image_image) also have "\ \ d \ {c * y..c * x}" by (auto simp add: field_simps 1) finally show "d \ (*) c ` {x..y} \ d \ {c * y..c * x}" . qed with \x \ y\ show ?thesis by auto qed (simp add: mult_left_mono_neg) qed lemma image_mult_atLeastAtMost_if': "(\x. x * c) ` {x..y} = (if x \ y then if c > 0 then {x * c .. y * c} else {y * c .. x * c} else {})" using image_mult_atLeastAtMost_if [of c x y] by (auto simp add: ac_simps) lemma image_affinity_atLeastAtMost: "((\x. m * x + c) ` {a..b}) = (if {a..b} = {} then {} else if 0 \ m then {m * a + c .. m * b + c} else {m * b + c .. m * a + c})" proof - have *: "(\x. m * x + c) = ((\x. x + c) \ (*) m)" by (simp add: fun_eq_iff) show ?thesis by (simp only: * image_comp [symmetric] image_mult_atLeastAtMost_if) (auto simp add: mult_le_cancel_left) qed lemma image_affinity_atLeastAtMost_diff: "((\x. m*x - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {m*a - c .. m*b - c} else {m*b - c .. m*a - c})" using image_affinity_atLeastAtMost [of m "-c" a b] by simp lemma image_affinity_atLeastAtMost_div: "((\x. x/m + c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m + c .. b/m + c} else {b/m + c .. a/m + c})" using image_affinity_atLeastAtMost [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) lemma image_affinity_atLeastAtMost_div_diff: "((\x. x/m - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m - c .. b/m - c} else {b/m - c .. a/m - c})" using image_affinity_atLeastAtMost_diff [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) end lemma atLeast1_lessThan_eq_remove0: "{Suc 0..x. x + (l::int)) ` {0..i. i - c) ` {x ..< y} = (if c < y then {x - c ..< y - c} else if x < y then {0} else {})" (is "_ = ?right") proof safe fix a assume a: "a \ ?right" show "a \ (\i. i - c) ` {x ..< y}" proof cases assume "c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ "a + c"]) next assume "\ c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ x] split: if_split_asm) qed qed auto lemma image_int_atLeastLessThan: "int ` {a..Finiteness\ lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..A bounded set of natural numbers is finite.\ lemma bounded_nat_set_is_finite: "(\i\N. i < (n::nat)) \ finite N" by (rule finite_subset [OF _ finite_lessThan]) auto text \A set of natural numbers is finite iff it is bounded.\ lemma finite_nat_set_iff_bounded: "finite(N::nat set) = (\m. \n\N. n?F\, simplified less_Suc_eq_le[symmetric]] by blast next assume ?B show ?F using \?B\ by(blast intro:bounded_nat_set_is_finite) qed lemma finite_nat_set_iff_bounded_le: "finite(N::nat set) = (\m. \n\N. n\m)" unfolding finite_nat_set_iff_bounded by (blast dest:less_imp_le_nat le_imp_less_Suc) lemma finite_less_ub: "!!f::nat=>nat. (!!n. n \ f n) ==> finite {n. f n \ u}" by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) lemma bounded_Max_nat: fixes P :: "nat \ bool" assumes x: "P x" and M: "\x. P x \ x \ M" obtains m where "P m" "\x. P x \ x \ m" proof - have "finite {x. P x}" using M finite_nat_set_iff_bounded_le by auto then have "Max {x. P x} \ {x. P x}" using Max_in x by auto then show ?thesis by (simp add: \finite {x. P x}\ that) qed text\Any subset of an interval of natural numbers the size of the subset is exactly that interval.\ lemma subset_card_intvl_is_intvl: assumes "A \ {k.. A" by auto with insert have "A \ {k..Proving Inclusions and Equalities between Unions\ lemma UN_le_eq_Un0: "(\i\n::nat. M i) = (\i\{1..n}. M i) \ M 0" (is "?A = ?B") proof show "?A \ ?B" proof fix x assume "x \ ?A" then obtain i where i: "i\n" "x \ M i" by auto show "x \ ?B" proof(cases i) case 0 with i show ?thesis by simp next case (Suc j) with i show ?thesis by auto qed qed next show "?B \ ?A" by fastforce qed lemma UN_le_add_shift: "(\i\n::nat. M(i+k)) = (\i\{k..n+k}. M i)" (is "?A = ?B") proof show "?A \ ?B" by fastforce next show "?B \ ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k..n+k}" "x \ M(i)" by auto hence "i-k\n \ x \ M((i-k)+k)" by auto thus "x \ ?A" by blast qed qed lemma UN_le_add_shift_strict: "(\ii\{k.. ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k.. M(i)" by auto then have "i - k < n \ x \ M((i-k) + k)" by auto then show "x \ ?A" using UN_le_add_shift by blast qed qed (fastforce) lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" by (auto simp add: atLeast0LessThan) lemma UN_finite_subset: "(\n::nat. (\i\{0.. C) \ (\n. A n) \ C" by (subst UN_UN_finite_eq [symmetric]) blast lemma UN_finite2_subset: assumes "\n::nat. (\i\{0.. (\i\{0..n. A n) \ (\n. B n)" proof (rule UN_finite_subset, rule) fix n and a from assms have "(\i\{0.. (\i\{0.. (\i\{0.. (\i\{0.. (\i. B i)" by (auto simp add: UN_UN_finite_eq) qed lemma UN_finite2_eq: "(\n::nat. (\i\{0..i\{0.. (\n. A n) = (\n. B n)" apply (rule subset_antisym [OF UN_finite_subset UN_finite2_subset]) apply auto apply (force simp add: atLeastLessThan_add_Un [of 0])+ done subsubsection \Cardinality\ lemma card_lessThan [simp]: "card {..x. x + l) ` {.. {0.. {0..n}" shows "finite N" using assms finite_atLeastAtMost by (rule finite_subset) lemma ex_bij_betw_nat_finite: "finite M \ \h. bij_betw h {0.. \h. bij_betw h M {0.. finite B \ card A = card B \ \h. bij_betw h A B" apply(drule ex_bij_betw_finite_nat) apply(drule ex_bij_betw_nat_finite) apply(auto intro!:bij_betw_trans) done lemma ex_bij_betw_nat_finite_1: "finite M \ \h. bij_betw h {1 .. card M} M" by (rule finite_same_card_bij) auto lemma bij_betw_iff_card: assumes "finite A" "finite B" shows "(\f. bij_betw f A B) \ (card A = card B)" proof assume "card A = card B" moreover obtain f where "bij_betw f A {0 ..< card A}" using assms ex_bij_betw_finite_nat by blast moreover obtain g where "bij_betw g {0 ..< card B} B" using assms ex_bij_betw_nat_finite by blast ultimately have "bij_betw (g \ f) A B" by (auto simp: bij_betw_trans) thus "(\f. bij_betw f A B)" by blast qed (auto simp: bij_betw_same_card) lemma subset_eq_atLeast0_lessThan_card: fixes n :: nat assumes "N \ {0.. n" proof - from assms finite_lessThan have "card N \ card {0..Relational version of @{thm [source] card_inj_on_le}:\ lemma card_le_if_inj_on_rel: assumes "finite B" "\a. a \ A \ \b. b\B \ r a b" "\a1 a2 b. \ a1 \ A; a2 \ A; b \ B; r a1 b; r a2 b \ \ a1 = a2" shows "card A \ card B" proof - let ?P = "\a b. b \ B \ r a b" let ?f = "\a. SOME b. ?P a b" have 1: "?f ` A \ B" by (auto intro: someI2_ex[OF assms(2)]) have "inj_on ?f A" proof (auto simp: inj_on_def) fix a1 a2 assume asms: "a1 \ A" "a2 \ A" "?f a1 = ?f a2" have 0: "?f a1 \ B" using "1" \a1 \ A\ by blast have 1: "r a1 (?f a1)" using someI_ex[OF assms(2)[OF \a1 \ A\]] by blast have 2: "r a2 (?f a1)" using someI_ex[OF assms(2)[OF \a2 \ A\]] asms(3) by auto show "a1 = a2" using assms(3)[OF asms(1,2) 0 1 2] . qed with 1 show ?thesis using card_inj_on_le[of ?f A B] assms(1) by simp qed lemma inj_on_funpow_least: \<^marker>\contributor \Lars Noschinski\\ \inj_on (\k. (f ^^ k) s) {0.. if \(f ^^ n) s = s\ \\m. 0 < m \ m < n \ (f ^^ m) s \ s\ proof - { fix k l assume A: "k < n" "l < n" "k \ l" "(f ^^ k) s = (f ^^ l) s" define k' l' where "k' = min k l" and "l' = max k l" with A have A': "k' < l'" "(f ^^ k') s = (f ^^ l') s" "l' < n" by (auto simp: min_def max_def) have "s = (f ^^ ((n - l') + l')) s" using that \l' < n\ by simp also have "\ = (f ^^ (n - l')) ((f ^^ l') s)" by (simp add: funpow_add) also have "(f ^^ l') s = (f ^^ k') s" by (simp add: A') also have "(f ^^ (n - l')) \ = (f ^^ (n - l' + k')) s" by (simp add: funpow_add) finally have "(f ^^ (n - l' + k')) s = s" by simp moreover have "n - l' + k' < n" "0 < n - l' + k'"using A' by linarith+ ultimately have False using that(2) by auto } then show ?thesis by (intro inj_onI) auto qed subsection \Intervals of integers\ lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..Finiteness\ lemma image_atLeastZeroLessThan_int: "0 \ u ==> {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int) qed auto lemma finite_atLeastLessThan_int [iff]: "finite {l..Cardinality\ lemma card_atLeastZeroLessThan_int: "card {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int card_image inj_on_def) qed auto lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" proof - have "{k. P k \ k < i} \ {.. M" shows "card {k \ M. k < Suc i} \ 0" proof - from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) qed lemma card_less_Suc2: assumes "0 \ M" shows "card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" proof - have *: "\j \ M; j < Suc i\ \ j - Suc 0 < i \ Suc (j - Suc 0) \ M \ Suc 0 \ j" for j by (cases j) (use assms in auto) show ?thesis proof (rule card_bij_eq) show "inj_on Suc {k. Suc k \ M \ k < i}" by force show "inj_on (\x. x - Suc 0) {k \ M. k < Suc i}" by (rule inj_on_diff_nat) (use * in blast) qed (use * in auto) qed lemma card_less_Suc: assumes "0 \ M" shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" proof - have "Suc (card {k. Suc k \ M \ k < i}) = Suc (card {k. Suc k \ M - {0} \ k < i})" by simp also have "\ = Suc (card {k \ M - {0}. k < Suc i})" apply (subst card_less_Suc2) using assms by auto also have "\ = Suc (card ({k \ M. k < Suc i} - {0}))" by (force intro: arg_cong [where f=card]) also have "\ = card (insert 0 ({k \ M. k < Suc i} - {0}))" by (simp add: card.insert_remove) also have "... = card {k \ M. k < Suc i}" using assms by (force simp add: intro: arg_cong [where f=card]) finally show ?thesis. qed lemma card_le_Suc_Max: "finite S \ card S \ Suc (Max S)" proof (rule classical) assume "finite S" and "\ Suc (Max S) \ card S" then have "Suc (Max S) < card S" by simp - with `finite S` have "S \ {0..Max S}" + with \finite S\ have "S \ {0..Max S}" by auto hence "card S \ card {0..Max S}" by (intro card_mono; auto) thus "card S \ Suc (Max S)" by simp qed subsection \Lemmas useful with the summation operator sum\ text \For examples, see Algebra/poly/UnivPoly2.thy\ subsubsection \Disjoint Unions\ text \Singletons and open intervals\ lemma ivl_disj_un_singleton: "{l::'a::linorder} Un {l<..} = {l..}" "{.. {l} Un {l<.. {l<.. u ==> {l} Un {l<..u} = {l..u}" "(l::'a::linorder) \ u ==> {l..One- and two-sided intervals\ lemma ivl_disj_un_one: "(l::'a::linorder) < u ==> {..l} Un {l<.. u ==> {.. u ==> {..l} Un {l<..u} = {..u}" "(l::'a::linorder) \ u ==> {.. u ==> {l<..u} Un {u<..} = {l<..}" "(l::'a::linorder) < u ==> {l<.. u ==> {l..u} Un {u<..} = {l..}" "(l::'a::linorder) \ u ==> {l..Two- and two-sided intervals\ lemma ivl_disj_un_two: "[| (l::'a::linorder) < m; m \ u |] ==> {l<.. m; m < u |] ==> {l<..m} Un {m<.. m; m \ u |] ==> {l.. m; m < u |] ==> {l..m} Un {m<.. u |] ==> {l<.. m; m \ u |] ==> {l<..m} Un {m<..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l.. m; m \ u |] ==> {l..m} Un {m<..u} = {l..u}" by auto lemma ivl_disj_un_two_touch: "[| (l::'a::linorder) < m; m < u |] ==> {l<..m} Un {m.. m; m < u |] ==> {l..m} Un {m.. u |] ==> {l<..m} Un {m..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l..m} Un {m..u} = {l..u}" by auto lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ivl_disj_un_two_touch subsubsection \Disjoint Intersections\ text \One- and two-sided intervals\ lemma ivl_disj_int_one: "{..l::'a::order} Int {l<..Two- and two-sided intervals\ lemma ivl_disj_int_two: "{l::'a::order<..Some Differences\ lemma ivl_diff[simp]: "i \ n \ {i..Some Subset Conditions\ lemma ivl_subset [simp]: "({i.. {m.. i \ m \ i \ j \ (n::'a::linorder))" using linorder_class.le_less_linear[of i n] apply (auto simp: linorder_not_le) apply (force intro: leI)+ done lemma obtain_subset_with_card_n: assumes "n \ card S" obtains T where "T \ S" "card T = n" "finite T" proof - obtain n' where "card S = n + n'" by (metis assms le_add_diff_inverse) with that show thesis proof (induct n' arbitrary: S) case 0 then show ?case by (cases "finite S") auto next case Suc then show ?case by (simp add: card_Suc_eq) (metis subset_insertI2) qed qed subsection \Generic big monoid operation over intervals\ context semiring_char_0 begin lemma inj_on_of_nat [simp]: "inj_on of_nat N" by rule simp lemma bij_betw_of_nat [simp]: "bij_betw of_nat N A \ of_nat ` N = A" by (simp add: bij_betw_def) end context comm_monoid_set begin lemma atLeastLessThan_reindex: "F g {h m.. h) {m.. h) {m..n}" if "bij_betw h {m..n} {h m..h n}" for m n ::nat proof - from that have "inj_on h {m..n}" and "h ` {m..n} = {h m..h n}" by (simp_all add: bij_betw_def) then show ?thesis using reindex [of h "{m..n}" g] by simp qed lemma atLeastLessThan_shift_bounds: "F g {m + k.. plus k) {m.. plus k) {m..n}" for m n k :: nat using atLeastAtMost_reindex [of "plus k" m n g] by (simp add: ac_simps) lemma atLeast_Suc_lessThan_Suc_shift: "F g {Suc m.. Suc) {m.. Suc) {m..n}" using atLeastAtMost_shift_bounds [of _ _ 1] by (simp add: plus_1_eq_Suc) lemma atLeast_int_lessThan_int_shift: "F g {int m.. int) {m.. int) {m..n}" by (rule atLeastAtMost_reindex) (simp add: image_int_atLeastAtMost) lemma atLeast0_lessThan_Suc: "F g {0..* g n" by (simp add: atLeast0_lessThan_Suc ac_simps) lemma atLeast0_atMost_Suc: "F g {0..Suc n} = F g {0..n} \<^bold>* g (Suc n)" by (simp add: atLeast0_atMost_Suc ac_simps) lemma atLeast0_lessThan_Suc_shift: "F g {0..* F (g \ Suc) {0..* F (g \ Suc) {0..n}" by (simp add: atLeast0_atMost_Suc_eq_insert_0 atLeast_Suc_atMost_Suc_shift) lemma atLeast_Suc_lessThan: "F g {m..* F g {Suc m..* F g {Suc m..n}" if "m \ n" proof - from that have "{m..n} = insert m {Suc m..n}" by auto then show ?thesis by simp qed lemma ivl_cong: "a = c \ b = d \ (\x. c \ x \ x < d \ g x = h x) \ F g {a.. plus m) {0.. n") simp_all lemma atLeastAtMost_shift_0: fixes m n p :: nat assumes "m \ n" shows "F g {m..n} = F (g \ plus m) {0..n - m}" using assms atLeastAtMost_shift_bounds [of g 0 m "n - m"] by simp lemma atLeastLessThan_concat: fixes m n p :: nat shows "m \ n \ n \ p \ F g {m..* F g {n..i. g (m + n - Suc i)) {n..i. g (m + n - i)) {n..m}" by (rule reindex_bij_witness [where i="\i. m + n - i" and j="\i. m + n - i"]) auto lemma atLeastLessThan_rev_at_least_Suc_atMost: "F g {n..i. g (m + n - i)) {Suc n..m}" unfolding atLeastLessThan_rev [of g n m] by (cases m) (simp_all add: atLeast_Suc_atMost_Suc_shift atLeastLessThanSuc_atLeastAtMost) end subsection \Summation indexed over intervals\ syntax (ASCII) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) syntax (latex_sum output) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" == "CONST sum (\x. t) {a..b}" "\x=a..x. t) {a..i\n. t" == "CONST sum (\i. t) {..n}" "\ii. t) {..The above introduces some pretty alternative syntaxes for summation over intervals: \begin{center} \begin{tabular}{lll} Old & New & \LaTeX\\ @{term[source]"\x\{a..b}. e"} & \<^term>\\x=a..b. e\ & @{term[mode=latex_sum]"\x=a..b. e"}\\ @{term[source]"\x\{a..\\x=a.. & @{term[mode=latex_sum]"\x=a..x\{..b}. e"} & \<^term>\\x\b. e\ & @{term[mode=latex_sum]"\x\b. e"}\\ @{term[source]"\x\{..\\x & @{term[mode=latex_sum]"\xlatex_sum\ (e.g.\ via \mode = latex_sum\ in antiquotations). It is not the default \LaTeX\ output because it only works well with italic-style formulae, not tt-style. Note that for uniformity on \<^typ>\nat\ it is better to use \<^term>\\x::nat=0.. rather than \\x: \sum\ may not provide all lemmas available for \<^term>\{m.. also in the special form for \<^term>\{...\ text\This congruence rule should be used for sums over intervals as the standard theorem @{text[source]sum.cong} does not work well with the simplifier who adds the unsimplified premise \<^term>\x\B\ to the context.\ context comm_monoid_set begin lemma zero_middle: assumes "1 \ p" "k \ p" shows "F (\j. if j < k then g j else if j = k then \<^bold>1 else h (j - Suc 0)) {..p} = F (\j. if j < k then g j else h j) {..p - Suc 0}" (is "?lhs = ?rhs") proof - have [simp]: "{..p - Suc 0} \ {j. j < k} = {.. - {j. j < k} = {k..p - Suc 0}" using assms by auto have "?lhs = F g {..* F (\j. if j = k then \<^bold>1 else h (j - Suc 0)) {k..p}" using union_disjoint [of "{.. = F g {..* F (\j. h (j - Suc 0)) {Suc k..p}" by (simp add: atLeast_Suc_atMost [of k p] assms) also have "\ = F g {..* F h {k .. p - Suc 0}" using reindex [of Suc "{k..p - Suc 0}"] assms by simp also have "\ = ?rhs" by (simp add: If_cases) finally show ?thesis . qed lemma atMost_Suc [simp]: "F g {..Suc n} = F g {..n} \<^bold>* g (Suc n)" by (simp add: atMost_Suc ac_simps) lemma lessThan_Suc [simp]: "F g {..* g n" by (simp add: lessThan_Suc ac_simps) lemma cl_ivl_Suc [simp]: "F g {m..Suc n} = (if Suc n < m then \<^bold>1 else F g {m..n} \<^bold>* g(Suc n))" by (auto simp: ac_simps atLeastAtMostSuc_conv) lemma op_ivl_Suc [simp]: "F g {m..1 else F g {m..* g(n))" by (auto simp: ac_simps atLeastLessThanSuc) lemma head: fixes n :: nat assumes mn: "m \ n" shows "F g {m..n} = g m \<^bold>* F g {m<..n}" (is "?lhs = ?rhs") proof - from mn have "{m..n} = {m} \ {m<..n}" by (auto intro: ivl_disj_un_singleton) hence "?lhs = F g ({m} \ {m<..n})" by (simp add: atLeast0LessThan) also have "\ = ?rhs" by simp finally show ?thesis . qed lemma last_plus: fixes n::nat shows "m \ n \ F g {m..n} = g n \<^bold>* F g {m..1 else F g {m..* g(n))" by (simp add: commute last_plus) lemma ub_add_nat: assumes "(m::nat) \ n + 1" shows "F g {m..n + p} = F g {m..n} \<^bold>* F g {n + 1..n + p}" proof- have "{m .. n+p} = {m..n} \ {n+1..n+p}" using \m \ n+1\ by auto thus ?thesis by (auto simp: ivl_disj_int union_disjoint atLeastSucAtMost_greaterThanAtMost) qed lemma nat_group: fixes k::nat shows "F (\m. F g {m * k ..< m*k + k}) {.. 0" by auto then show ?thesis by (induct n) (simp_all add: atLeastLessThan_concat add.commute atLeast0LessThan[symmetric]) qed auto lemma triangle_reindex: fixes n :: nat shows "F (\(i,j). g i j) {(i,j). i+j < n} = F (\k. F (\i. g i (k - i)) {..k}) {..(i,j). g i j) {(i,j). i+j \ n} = F (\k. F (\i. g i (k - i)) {..k}) {..n}" using triangle_reindex [of g "Suc n"] by (simp only: Nat.less_Suc_eq_le lessThan_Suc_atMost) lemma nat_diff_reindex: "F (\i. g (n - Suc i)) {..i. g(i + k)){m..i. g(i + k)){m..n::nat}" by (rule reindex_bij_witness[where i="\i. i + k" and j="\i. i - k"]) auto corollary shift_bounds_cl_Suc_ivl: "F g {Suc m..Suc n} = F (\i. g(Suc i)){m..n}" by (simp add: shift_bounds_cl_nat_ivl[where k="Suc 0", simplified]) corollary Suc_reindex_ivl: "m \ n \ F g {m..n} \<^bold>* g (Suc n) = g m \<^bold>* F (\i. g (Suc i)) {m..n}" by (simp add: assoc atLeast_Suc_atMost flip: shift_bounds_cl_Suc_ivl) corollary shift_bounds_Suc_ivl: "F g {Suc m..i. g(Suc i)){m..* F (\i. g (Suc i)) {..n}" proof (induct n) case 0 show ?case by simp next case (Suc n) note IH = this have "F g {..Suc (Suc n)} = F g {..Suc n} \<^bold>* g (Suc (Suc n))" by (rule atMost_Suc) also have "F g {..Suc n} = g 0 \<^bold>* F (\i. g (Suc i)) {..n}" by (rule IH) also have "g 0 \<^bold>* F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = g 0 \<^bold>* (F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)))" by (rule assoc) also have "F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = F (\i. g (Suc i)) {..Suc n}" by (rule atMost_Suc [symmetric]) finally show ?case . qed lemma lessThan_Suc_shift: "F g {..* F (\i. g (Suc i)) {..* F (\i. g (Suc i)) {..i. F (\j. a i j) {0..j. F (\i. a i j) {Suc j..n}) {0..i. F (\j. a i j) {..j. F (\i. a i j) {Suc j..n}) {..k. g (Suc k)) {.. = F (\k. g (Suc k)) {.. b \ F g {a..* g b" by (simp add: atLeastLessThanSuc commute) lemma nat_ivl_Suc': assumes "m \ Suc n" shows "F g {m..Suc n} = g (Suc n) \<^bold>* F g {m..n}" proof - from assms have "{m..Suc n} = insert (Suc n) {m..n}" by auto also have "F g \ = g (Suc n) \<^bold>* F g {m..n}" by simp finally show ?thesis . qed lemma in_pairs: "F g {2*m..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {m..n}" proof (induction n) case 0 show ?case by (cases "m=0") auto next case (Suc n) then show ?case by (auto simp: assoc split: if_split_asm) qed lemma in_pairs_0: "F g {..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {..n}" using in_pairs [of _ 0 n] by (simp add: atLeast0AtMost) end lemma card_sum_le_nat_sum: "\ {0.. \ S" proof (cases "finite S") case True then show ?thesis proof (induction "card S" arbitrary: S) case (Suc x) then have "Max S \ x" using card_le_Suc_Max by fastforce let ?S' = "S - {Max S}" from Suc have "Max S \ S" by (auto intro: Max_in) hence cards: "card S = Suc (card ?S')" - using `finite S` by (intro card.remove; auto) + using \finite S\ by (intro card.remove; auto) hence "\ {0.. \ ?S'" using Suc by (intro Suc; auto) hence "\ {0.. \ ?S' + Max S" - using `Max S \ x` by simp + using \Max S \ x\ by simp also have "... = \ S" - using sum.remove[OF `finite S` `Max S \ S`, where g="\x. x"] + using sum.remove[OF \finite S\ \Max S \ S\, where g="\x. x"] by simp finally show ?case using cards Suc by auto qed simp qed simp lemma sum_natinterval_diff: fixes f:: "nat \ ('a::ab_group_add)" shows "sum (\k. f k - f(k + 1)) {(m::nat) .. n} = (if m \ n then f m - f(n + 1) else 0)" by (induct n, auto simp add: algebra_simps not_le le_Suc_eq) lemma sum_diff_nat_ivl: fixes f :: "nat \ 'a::ab_group_add" shows "\ m \ n; n \ p \ \ sum f {m..x. Q x \ P x \ (\xxxShifting bounds\ context comm_monoid_add begin context fixes f :: "nat \ 'a" assumes "f 0 = 0" begin lemma sum_shift_lb_Suc0_0_upt: "sum f {Suc 0..f 0 = 0\ by simp qed lemma sum_shift_lb_Suc0_0: "sum f {Suc 0..k} = sum f {0..k}" proof (cases k) case 0 with \f 0 = 0\ show ?thesis by simp next case (Suc k) moreover have "{0..Suc k} = insert 0 {Suc 0..Suc k}" by auto ultimately show ?thesis using \f 0 = 0\ by simp qed end end lemma sum_Suc_diff: fixes f :: "nat \ 'a::ab_group_add" assumes "m \ Suc n" shows "(\i = m..n. f(Suc i) - f i) = f (Suc n) - f m" using assms by (induct n) (auto simp: le_Suc_eq) lemma sum_Suc_diff': fixes f :: "nat \ 'a::ab_group_add" assumes "m \ n" shows "(\i = m..Telescoping\ lemma sum_telescope: fixes f::"nat \ 'a::ab_group_add" shows "sum (\i. f i - f (Suc i)) {.. i} = f 0 - f (Suc i)" by (induct i) simp_all lemma sum_telescope'': assumes "m \ n" shows "(\k\{Suc m..n}. f k - f (k - 1)) = f n - (f m :: 'a :: ab_group_add)" by (rule dec_induct[OF assms]) (simp_all add: algebra_simps) lemma sum_lessThan_telescope: "(\nnThe formula for geometric sums\ lemma sum_power2: "(\i=0.. 1" shows "(\i 0" by simp_all moreover have "(\iy \ 0\) ultimately show ?thesis by simp qed lemma diff_power_eq_sum: fixes y :: "'a::{comm_ring,monoid_mult}" shows "x ^ (Suc n) - y ^ (Suc n) = (x - y) * (\pppp \\COMPLEX_POLYFUN\ in HOL Light\ fixes x :: "'a::{comm_ring,monoid_mult}" shows "x^n - y^n = (x - y) * (\iiiii\n. x^i) = 1 - x^Suc n" by (simp only: one_diff_power_eq lessThan_Suc_atMost) lemma sum_power_shift: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(\i=m..n. x^i) = x^m * (\i\n-m. x^i)" proof - have "(\i=m..n. x^i) = x^m * (\i=m..n. x^(i-m))" by (simp add: sum_distrib_left power_add [symmetric]) also have "(\i=m..n. x^(i-m)) = (\i\n-m. x^i)" using \m \ n\ by (intro sum.reindex_bij_witness[where j="\i. i - m" and i="\i. i + m"]) auto finally show ?thesis . qed lemma sum_gp_multiplied: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(1 - x) * (\i=m..n. x^i) = x^m - x^Suc n" proof - have "(1 - x) * (\i=m..n. x^i) = x^m * (1 - x) * (\i\n-m. x^i)" by (metis mult.assoc mult.commute assms sum_power_shift) also have "... =x^m * (1 - x^Suc(n-m))" by (metis mult.assoc sum_gp_basic) also have "... = x^m - x^Suc n" using assms by (simp add: algebra_simps) (metis le_add_diff_inverse power_add) finally show ?thesis . qed lemma sum_gp: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..n. x^i) = (if n < m then 0 else if x = 1 then of_nat((n + 1) - m) else (x^m - x^Suc n) / (1 - x))" using sum_gp_multiplied [of m n x] apply auto by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) subsubsection\Geometric progressions\ lemma sum_gp0: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i\n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))" using sum_gp_basic[of x n] by (simp add: mult.commute field_split_simps) lemma sum_power_add: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\I. x^(m+i)) = x^m * (\i\I. x^i)" by (simp add: sum_distrib_left power_add) lemma sum_gp_offset: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..m+n. x^i) = (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" using sum_gp [of x m "m+n"] by (auto simp: power_add algebra_simps) lemma sum_gp_strict: fixes x :: "'a::{comm_ring,division_ring}" shows "(\iThe formulae for arithmetic sums\ context comm_semiring_1 begin lemma double_gauss_sum: "2 * (\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1)" by (induct n) (simp_all add: sum.atLeast0_atMost_Suc algebra_simps left_add_twice) lemma double_gauss_sum_from_Suc_0: "2 * (\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1)" proof - have "sum of_nat {Suc 0..n} = sum of_nat (insert 0 {Suc 0..n})" by simp also have "\ = sum of_nat {0..n}" by (cases n) (simp_all add: atLeast0_atMost_Suc_eq_insert_0) finally show ?thesis by (simp add: double_gauss_sum) qed lemma double_arith_series: "2 * (\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d)" proof - have "(\i = 0..n. a + of_nat i * d) = ((\i = 0..n. a) + (\i = 0..n. of_nat i * d))" by (rule sum.distrib) also have "\ = (of_nat (Suc n) * a + d * (\i = 0..n. of_nat i))" by (simp add: sum_distrib_left algebra_simps) finally show ?thesis by (simp add: algebra_simps double_gauss_sum left_add_twice) qed end context unique_euclidean_semiring_with_nat begin lemma gauss_sum: "(\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum [of n, symmetric] by simp lemma gauss_sum_from_Suc_0: "(\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum_from_Suc_0 [of n, symmetric] by simp lemma arith_series: "(\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d) div 2" using double_arith_series [of a d n, symmetric] by simp end lemma gauss_sum_nat: "\{0..n} = (n * Suc n) div 2" using gauss_sum [of n, where ?'a = nat] by simp lemma arith_series_nat: "(\i = 0..n. a + i * d) = Suc n * (2 * a + n * d) div 2" using arith_series [of a d n] by simp lemma Sum_Icc_int: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" if "m \ n" for m n :: int using that proof (induct i \ "nat (n - m)" arbitrary: m n) case 0 then have "m = n" by arith then show ?case by (simp add: algebra_simps mult_2 [symmetric]) next case (Suc i) have 0: "i = nat((n-1) - m)" "m \ n-1" using Suc(2,3) by arith+ have "\ {m..n} = \ {m..1+(n-1)}" by simp also have "\ = \ {m..n-1} + n" using \m \ n\ by(subst atLeastAtMostPlus1_int_conv) simp_all also have "\ = ((n-1)*(n-1+1) - m*(m-1)) div 2 + n" by(simp add: Suc(1)[OF 0]) also have "\ = ((n-1)*(n-1+1) - m*(m-1) + 2*n) div 2" by simp also have "\ = (n*(n+1) - m*(m-1)) div 2" by (simp add: algebra_simps mult_2_right) finally show ?case . qed lemma Sum_Icc_nat: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" for m n :: nat proof (cases "m \ n") case True then have *: "m * (m - 1) \ n * (n + 1)" by (meson diff_le_self order_trans le_add1 mult_le_mono) have "int (\{m..n}) = (\{int m..int n})" by (simp add: sum.atLeast_int_atMost_int_shift) also have "\ = (int n * (int n + 1) - int m * (int m - 1)) div 2" using \m \ n\ by (simp add: Sum_Icc_int) also have "\ = int ((n * (n + 1) - m * (m - 1)) div 2)" using le_square * by (simp add: algebra_simps of_nat_div of_nat_diff) finally show ?thesis by (simp only: of_nat_eq_iff) next case False then show ?thesis by (auto dest: less_imp_Suc_add simp add: not_le algebra_simps) qed lemma Sum_Ico_nat: "\{m..Division remainder\ lemma range_mod: fixes n :: nat assumes "n > 0" shows "range (\m. m mod n) = {0.. ?A \ m \ ?B" proof assume "m \ ?A" with assms show "m \ ?B" by auto next assume "m \ ?B" moreover have "m mod n \ ?A" by (rule rangeI) ultimately show "m \ ?A" by simp qed qed subsection \Products indexed over intervals\ syntax (ASCII) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) syntax (latex_prod output) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" \ "CONST prod (\x. t) {a..b}" "\x=a.. "CONST prod (\x. t) {a..i\n. t" \ "CONST prod (\i. t) {..n}" "\i "CONST prod (\i. t) {..{int i..int (i+j)}" by (induct j) (auto simp add: atLeastAtMostSuc_conv atLeastAtMostPlus1_int_conv) lemma prod_int_eq: "prod int {i..j} = \{int i..int j}" proof (cases "i \ j") case True then show ?thesis by (metis le_iff_add prod_int_plus_eq) next case False then show ?thesis by auto qed subsection \Efficient folding over intervals\ function fold_atLeastAtMost_nat where [simp del]: "fold_atLeastAtMost_nat f a (b::nat) acc = (if a > b then acc else fold_atLeastAtMost_nat f (a+1) b (f a acc))" by pat_completeness auto termination by (relation "measure (\(_,a,b,_). Suc b - a)") auto lemma fold_atLeastAtMost_nat: assumes "comp_fun_commute f" shows "fold_atLeastAtMost_nat f a b acc = Finite_Set.fold f acc {a..b}" using assms proof (induction f a b acc rule: fold_atLeastAtMost_nat.induct, goal_cases) case (1 f a b acc) interpret comp_fun_commute f by fact show ?case proof (cases "a > b") case True thus ?thesis by (subst fold_atLeastAtMost_nat.simps) auto next case False with 1 show ?thesis by (subst fold_atLeastAtMost_nat.simps) (auto simp: atLeastAtMost_insertL[symmetric] fold_fun_left_comm) qed qed lemma sum_atLeastAtMost_code: "sum f {a..b} = fold_atLeastAtMost_nat (\a acc. f a + acc) a b 0" proof - have "comp_fun_commute (\a. (+) (f a))" by unfold_locales (auto simp: o_def add_ac) thus ?thesis by (simp add: sum.eq_fold fold_atLeastAtMost_nat o_def) qed lemma prod_atLeastAtMost_code: "prod f {a..b} = fold_atLeastAtMost_nat (\a acc. f a * acc) a b 1" proof - have "comp_fun_commute (\a. (*) (f a))" by unfold_locales (auto simp: o_def mult_ac) thus ?thesis by (simp add: prod.eq_fold fold_atLeastAtMost_nat o_def) qed (* TODO: Add support for folding over more kinds of intervals here *) end