diff -r 714fad33252e -r 83e4b6a5e7de src/HOL/Divides.thy --- a/src/HOL/Divides.thy Fri Aug 19 05:49:11 2022 +0000 +++ b/src/HOL/Divides.thy Fri Aug 19 05:49:12 2022 +0000 @@ -11,112 +11,8 @@ subsection \More on division\ -subsubsection \Splitting Rules for div and mod\ - -text\The proofs of the two lemmas below are essentially identical\ - -lemma split_pos_lemma: - "0 - P(n div k :: int)(n mod k) = (\i j. 0\j \ j n = k*i + j \ P i j)" - by auto - -lemma split_neg_lemma: - "k<0 \ - P(n div k :: int)(n mod k) = (\i j. k j\0 \ n = k*i + j \ P i j)" - by auto - -lemma split_zdiv: - \P (n div k) \ - (k = 0 \ P 0) \ - (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ P i)) \ - (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ P i))\ for n k :: int -proof (cases \k = 0\) - case True - then show ?thesis - by simp -next - case False - then have \k < 0 \ 0 < k\ - by auto - then show ?thesis - by (auto simp add: split_pos_lemma [of concl: "\x y. P x"] split_neg_lemma [of concl: "\x y. P x"]) -qed - -lemma split_zmod: - \P (n mod k) \ - (k = 0 \ P n) \ - (0 < k \ (\i j. 0 \ j \ j < k \ n = k * i + j \ P j)) \ - (k < 0 \ (\i j. k < j \ j \ 0 \ n = k * i + j \ P j))\ for n k :: int -proof (cases \k = 0\) - case True - then show ?thesis - by simp -next - case False - then have \k < 0 \ 0 < k\ - by auto - then show ?thesis - by (auto simp add: split_pos_lemma [of concl: "\x y. P y"] split_neg_lemma [of concl: "\x y. P y"]) -qed - -text \Enable (lin)arith to deal with \<^const>\divide\ and \<^const>\modulo\ - when these are applied to some constant that is of the form - \<^term>\numeral k\:\ -declare split_zdiv [of _ _ \numeral n\, linarith_split] for n -declare split_zdiv [of _ _ \- numeral n\, linarith_split] for n -declare split_zmod [of _ _ \numeral n\, linarith_split] for n -declare split_zmod [of _ _ \- numeral n\, linarith_split] for n - -lemma half_nonnegative_int_iff [simp]: - \k div 2 \ 0 \ k \ 0\ for k :: int - by auto - -lemma half_negative_int_iff [simp]: - \k div 2 < 0 \ k < 0\ for k :: int - by auto - -lemma zdiv_eq_0_iff: - "i div k = 0 \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" (is "?L = ?R") - for i k :: int -proof - assume ?L - moreover have "?L \ ?R" - by (rule split_zdiv [THEN iffD2]) simp - ultimately show ?R - by blast -next - assume ?R then show ?L - by auto -qed - -lemma zmod_trivial_iff: - fixes i k :: int - shows "i mod k = i \ k = 0 \ 0 \ i \ i < k \ i \ 0 \ k < i" -proof - - have "i mod k = i \ i div k = 0" - using div_mult_mod_eq [of i k] by safe auto - with zdiv_eq_0_iff - show ?thesis - by simp -qed - - subsubsection \Monotonicity in the First Argument (Dividend)\ -inductive eucl_rel_int :: "int \ int \ int \ int \ bool" - where eucl_rel_int_by0: "eucl_rel_int k 0 (0, k)" - | eucl_rel_int_dividesI: "l \ 0 \ k = q * l \ eucl_rel_int k l (q, 0)" - | eucl_rel_int_remainderI: "sgn r = sgn l \ \r\ < \l\ - \ k = q * l + r \ eucl_rel_int k l (q, r)" - -lemma eucl_rel_int_iff: - "eucl_rel_int k l (q, r) \ - k = l * q + r \ - (if 0 < l then 0 \ r \ r < l else if l < 0 then l < r \ r \ 0 else q = 0)" - by (cases "r = 0") - (auto elim!: eucl_rel_int.cases intro: eucl_rel_int_by0 eucl_rel_int_dividesI eucl_rel_int_remainderI - simp add: ac_simps sgn_1_pos sgn_1_neg) - lemma unique_quotient_lemma: assumes "b * q' + r' \ b * q + r" "0 \ r'" "r' < b" "r < b" shows "q' \ (q::int)" proof - @@ -134,61 +30,6 @@ "b * q' + r' \ b*q + r \ r \ 0 \ b < r \ b < r' \ q \ (q'::int)" using unique_quotient_lemma[where b = "-b" and r = "-r'" and r'="-r"] by auto -lemma unique_quotient: - "eucl_rel_int a b (q, r) \ eucl_rel_int a b (q', r') \ q = q'" - apply (rule order_antisym) - apply (simp_all add: eucl_rel_int_iff linorder_neq_iff split: if_split_asm) - apply (blast intro: order_eq_refl [THEN unique_quotient_lemma] order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ - done - -lemma unique_remainder: - assumes "eucl_rel_int a b (q, r)" - and "eucl_rel_int a b (q', r')" - shows "r = r'" -proof - - have "q = q'" - using assms by (blast intro: unique_quotient) - then show "r = r'" - using assms by (simp add: eucl_rel_int_iff) -qed - -lemma eucl_rel_int: - "eucl_rel_int k l (k div l, k mod l)" -proof (cases k rule: int_cases3) - case zero - then show ?thesis - by (simp add: eucl_rel_int_iff divide_int_def modulo_int_def) -next - case (pos n) - then show ?thesis - using div_mult_mod_eq [of n] - by (cases l rule: int_cases3) - (auto simp del: of_nat_mult of_nat_add - simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps - eucl_rel_int_iff divide_int_def modulo_int_def) -next - case (neg n) - then show ?thesis - using div_mult_mod_eq [of n] - by (cases l rule: int_cases3) - (auto simp del: of_nat_mult of_nat_add - simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps - eucl_rel_int_iff divide_int_def modulo_int_def) -qed - -lemma divmod_int_unique: - assumes "eucl_rel_int k l (q, r)" - shows div_int_unique: "k div l = q" and mod_int_unique: "k mod l = r" - using assms eucl_rel_int [of k l] - using unique_quotient [of k l] unique_remainder [of k l] - by auto - -lemma zminus1_lemma: - "eucl_rel_int a b (q, r) ==> b \ 0 - ==> eucl_rel_int (-a) b (if r=0 then -q else -q - 1, - if r=0 then 0 else b-r)" -by (force simp add: eucl_rel_int_iff right_diff_distrib) - lemma zdiv_mono1: \a div b \ a' div b\ if \a \ a'\ \0 < b\ @@ -272,6 +113,72 @@ by simp qed (use assms in auto) + +subsubsection \Computing \div\ and \mod\ with shifting\ + +inductive eucl_rel_int :: "int \ int \ int \ int \ bool" + where eucl_rel_int_by0: "eucl_rel_int k 0 (0, k)" + | eucl_rel_int_dividesI: "l \ 0 \ k = q * l \ eucl_rel_int k l (q, 0)" + | eucl_rel_int_remainderI: "sgn r = sgn l \ \r\ < \l\ + \ k = q * l + r \ eucl_rel_int k l (q, r)" + +lemma eucl_rel_int_iff: + "eucl_rel_int k l (q, r) \ + k = l * q + r \ + (if 0 < l then 0 \ r \ r < l else if l < 0 then l < r \ r \ 0 else q = 0)" + by (cases "r = 0") + (auto elim!: eucl_rel_int.cases intro: eucl_rel_int_by0 eucl_rel_int_dividesI eucl_rel_int_remainderI + simp add: ac_simps sgn_1_pos sgn_1_neg) + +lemma unique_quotient: + "eucl_rel_int a b (q, r) \ eucl_rel_int a b (q', r') \ q = q'" + apply (rule order_antisym) + apply (simp_all add: eucl_rel_int_iff linorder_neq_iff split: if_split_asm) + apply (blast intro: order_eq_refl [THEN unique_quotient_lemma] order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ + done + +lemma unique_remainder: + assumes "eucl_rel_int a b (q, r)" + and "eucl_rel_int a b (q', r')" + shows "r = r'" +proof - + have "q = q'" + using assms by (blast intro: unique_quotient) + then show "r = r'" + using assms by (simp add: eucl_rel_int_iff) +qed + +lemma eucl_rel_int: + "eucl_rel_int k l (k div l, k mod l)" +proof (cases k rule: int_cases3) + case zero + then show ?thesis + by (simp add: eucl_rel_int_iff divide_int_def modulo_int_def) +next + case (pos n) + then show ?thesis + using div_mult_mod_eq [of n] + by (cases l rule: int_cases3) + (auto simp del: of_nat_mult of_nat_add + simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps + eucl_rel_int_iff divide_int_def modulo_int_def) +next + case (neg n) + then show ?thesis + using div_mult_mod_eq [of n] + by (cases l rule: int_cases3) + (auto simp del: of_nat_mult of_nat_add + simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps + eucl_rel_int_iff divide_int_def modulo_int_def) +qed + +lemma divmod_int_unique: + assumes "eucl_rel_int k l (q, r)" + shows div_int_unique: "k div l = q" and mod_int_unique: "k mod l = r" + using assms eucl_rel_int [of k l] + using unique_quotient [of k l] unique_remainder [of k l] + by auto + lemma div_pos_geq: fixes k l :: int assumes "0 < l" and "l \ k" @@ -292,9 +199,6 @@ with assms show ?thesis by simp qed - -subsubsection \Computing \div\ and \mod\ with shifting\ - lemma pos_eucl_rel_int_mult_2: assumes "0 \ b" assumes "eucl_rel_int a b (q, r)"