src/HOL/Divides.thy
changeset 75936 d2e6a1342c90
parent 75883 d7e0b6620c07
child 75937 02b18f59f903
equal deleted inserted replaced
75935:06eb4d0031e3 75936:d2e6a1342c90
   544 subsection \<open>Numeral division with a pragmatic type class\<close>
   544 subsection \<open>Numeral division with a pragmatic type class\<close>
   545 
   545 
   546 text \<open>
   546 text \<open>
   547   The following type class contains everything necessary to formulate
   547   The following type class contains everything necessary to formulate
   548   a division algorithm in ring structures with numerals, restricted
   548   a division algorithm in ring structures with numerals, restricted
   549   to its positive segments.  This is its primary motivation, and it
   549   to its positive segments.
   550   could surely be formulated using a more fine-grained, more algebraic
       
   551   and less technical class hierarchy.
       
   552 \<close>
   550 \<close>
   553 
   551 
   554 class unique_euclidean_semiring_numeral = unique_euclidean_semiring_with_nat + linordered_semidom +
   552 class unique_euclidean_semiring_with_nat_division = unique_euclidean_semiring_with_nat +
   555   assumes div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0"
   553   fixes divmod :: \<open>num \<Rightarrow> num \<Rightarrow> 'a \<times> 'a\<close>
   556     and mod_less: " 0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a mod b = a"
   554     and divmod_step :: \<open>'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<times> 'a\<close> \<comment> \<open>
   557     and div_positive: "0 < b \<Longrightarrow> b \<le> a \<Longrightarrow> a div b > 0"
   555       These are conceptually definitions but force generated code
   558     and mod_less_eq_dividend: "0 \<le> a \<Longrightarrow> a mod b \<le> a"
   556       to be monomorphic wrt. particular instances of this class which
   559     and pos_mod_bound: "0 < b \<Longrightarrow> a mod b < b"
   557       yields a significant speedup.\<close>
   560     and pos_mod_sign: "0 < b \<Longrightarrow> 0 \<le> a mod b"
   558   assumes divmod_def: \<open>divmod m n = (numeral m div numeral n, numeral m mod numeral n)\<close>
   561     and mod_mult2_eq: "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b"
   559     and divmod_step_def [simp]: \<open>divmod_step l (q, r) =
   562     and div_mult2_eq: "0 \<le> c \<Longrightarrow> a div (b * c) = a div b div c"
   560       (if euclidean_size l \<le> euclidean_size r then (2 * q + 1, r - l)
   563   assumes discrete: "a < b \<longleftrightarrow> a + 1 \<le> b"
   561        else (2 * q, r))\<close> \<comment> \<open>
   564   fixes divmod :: "num \<Rightarrow> num \<Rightarrow> 'a \<times> 'a"
   562          This is a formulation of one step (referring to one digit position)
   565     and divmod_step :: "'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<times> 'a"
   563          in school-method division: compare the dividend at the current
   566   assumes divmod_def: "divmod m n = (numeral m div numeral n, numeral m mod numeral n)"
   564          digit position with the remainder from previous division steps
   567     and divmod_step_def: "divmod_step l qr = (let (q, r) = qr
   565          and evaluate accordingly.\<close>
   568     in if r \<ge> l then (2 * q + 1, r - l)
       
   569     else (2 * q, r))"
       
   570     \<comment> \<open>These are conceptually definitions but force generated code
       
   571     to be monomorphic wrt. particular instances of this class which
       
   572     yields a significant speedup.\<close>
       
   573 begin
   566 begin
   574 
   567 
   575 lemma divmod_digit_1:
       
   576   assumes "0 \<le> a" "0 < b" and "b \<le> a mod (2 * b)"
       
   577   shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P")
       
   578     and "a mod (2 * b) - b = a mod b" (is "?Q")
       
   579 proof -
       
   580   from assms mod_less_eq_dividend [of a "2 * b"] have "b \<le> a"
       
   581     by (auto intro: trans)
       
   582   with \<open>0 < b\<close> have "0 < a div b" by (auto intro: div_positive)
       
   583   then have [simp]: "1 \<le> a div b" by (simp add: discrete)
       
   584   with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound)
       
   585   define w where "w = a div b mod 2"
       
   586   then have w_exhaust: "w = 0 \<or> w = 1" by auto
       
   587   have mod_w: "a mod (2 * b) = a mod b + b * w"
       
   588     by (simp add: w_def mod_mult2_eq ac_simps)
       
   589   from assms w_exhaust have "w = 1"
       
   590     using mod_less by (auto simp add: mod_w)
       
   591   with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp
       
   592   have "2 * (a div (2 * b)) = a div b - w"
       
   593     by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps)
       
   594   with \<open>w = 1\<close> have div: "2 * (a div (2 * b)) = a div b - 1" by simp
       
   595   then show ?P and ?Q
       
   596     by (simp_all add: div mod add_implies_diff [symmetric])
       
   597 qed
       
   598 
       
   599 lemma divmod_digit_0:
       
   600   assumes "0 < b" and "a mod (2 * b) < b"
       
   601   shows "2 * (a div (2 * b)) = a div b" (is "?P")
       
   602     and "a mod (2 * b) = a mod b" (is "?Q")
       
   603 proof -
       
   604   define w where "w = a div b mod 2"
       
   605   then have w_exhaust: "w = 0 \<or> w = 1" by auto
       
   606   have mod_w: "a mod (2 * b) = a mod b + b * w"
       
   607     by (simp add: w_def mod_mult2_eq ac_simps)
       
   608   moreover have "b \<le> a mod b + b"
       
   609   proof -
       
   610     from \<open>0 < b\<close> pos_mod_sign have "0 \<le> a mod b" by blast
       
   611     then have "0 + b \<le> a mod b + b" by (rule add_right_mono)
       
   612     then show ?thesis by simp
       
   613   qed
       
   614   moreover note assms w_exhaust
       
   615   ultimately have "w = 0" by auto
       
   616   with mod_w have mod: "a mod (2 * b) = a mod b" by simp
       
   617   have "2 * (a div (2 * b)) = a div b - w"
       
   618     by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps)
       
   619   with \<open>w = 0\<close> have div: "2 * (a div (2 * b)) = a div b" by simp
       
   620   then show ?P and ?Q
       
   621     by (simp_all add: div mod)
       
   622 qed
       
   623 
       
   624 lemma mod_double_modulus:
       
   625   assumes "m > 0" "x \<ge> 0"
       
   626   shows   "x mod (2 * m) = x mod m \<or> x mod (2 * m) = x mod m + m"
       
   627 proof (cases "x mod (2 * m) < m")
       
   628   case True
       
   629   thus ?thesis using assms using divmod_digit_0(2)[of m x] by auto
       
   630 next
       
   631   case False
       
   632   hence *: "x mod (2 * m) - m = x mod m"
       
   633     using assms by (intro divmod_digit_1) auto
       
   634   hence "x mod (2 * m) = x mod m + m"
       
   635     by (subst * [symmetric], subst le_add_diff_inverse2) (use False in auto)
       
   636   thus ?thesis by simp
       
   637 qed
       
   638 
       
   639 lemma fst_divmod:
   568 lemma fst_divmod:
   640   "fst (divmod m n) = numeral m div numeral n"
   569   \<open>fst (divmod m n) = numeral m div numeral n\<close>
   641   by (simp add: divmod_def)
   570   by (simp add: divmod_def)
   642 
   571 
   643 lemma snd_divmod:
   572 lemma snd_divmod:
   644   "snd (divmod m n) = numeral m mod numeral n"
   573   \<open>snd (divmod m n) = numeral m mod numeral n\<close>
   645   by (simp add: divmod_def)
   574   by (simp add: divmod_def)
   646 
   575 
   647 text \<open>
   576 text \<open>
   648   This is a formulation of one step (referring to one digit position)
   577   Following a formulation of school-method division.
   649   in school-method division: compare the dividend at the current
       
   650   digit position with the remainder from previous division steps
       
   651   and evaluate accordingly.
       
   652 \<close>
       
   653 
       
   654 lemma divmod_step_eq [simp]:
       
   655   "divmod_step l (q, r) = (if l \<le> r
       
   656     then (2 * q + 1, r - l) else (2 * q, r))"
       
   657   by (simp add: divmod_step_def)
       
   658 
       
   659 text \<open>
       
   660   This is a formulation of school-method division.
       
   661   If the divisor is smaller than the dividend, terminate.
   578   If the divisor is smaller than the dividend, terminate.
   662   If not, shift the dividend to the right until termination
   579   If not, shift the dividend to the right until termination
   663   occurs and then reiterate single division steps in the
   580   occurs and then reiterate single division steps in the
   664   opposite direction.
   581   opposite direction.
   665 \<close>
   582 \<close>
   666 
   583 
   667 lemma divmod_divmod_step:
   584 lemma divmod_divmod_step:
   668   "divmod m n = (if m < n then (0, numeral m)
   585   \<open>divmod m n = (if m < n then (0, numeral m)
   669     else divmod_step (numeral n) (divmod m (Num.Bit0 n)))"
   586     else divmod_step (numeral n) (divmod m (Num.Bit0 n)))\<close>
   670 proof (cases "m < n")
   587 proof (cases \<open>m < n\<close>)
   671   case True then have "numeral m < numeral n" by simp
   588   case True
   672   then show ?thesis
   589   then show ?thesis
   673     by (simp add: prod_eq_iff div_less mod_less fst_divmod snd_divmod)
   590     by (simp add: prod_eq_iff fst_divmod snd_divmod flip: of_nat_numeral of_nat_div of_nat_mod)
   674 next
   591 next
   675   case False
   592   case False
   676   have "divmod m n =
   593   define r s t where \<open>r = (numeral m :: nat)\<close> \<open>s = (numeral n :: nat)\<close> \<open>t = 2 * s\<close>
   677     divmod_step (numeral n) (numeral m div (2 * numeral n),
   594   then have *: \<open>numeral m = of_nat r\<close> \<open>numeral n = of_nat s\<close> \<open>numeral (num.Bit0 n) = of_nat t\<close>
   678       numeral m mod (2 * numeral n))"
   595     and \<open>\<not> s \<le> r mod s\<close>
   679   proof (cases "numeral n \<le> numeral m mod (2 * numeral n)")
   596     by (simp_all add: not_le)
   680     case True
   597   have t: \<open>2 * (r div t) = r div s - r div s mod 2\<close>
   681     with divmod_step_eq
   598     \<open>r mod t = s * (r div s mod 2) + r mod s\<close>
   682       have "divmod_step (numeral n) (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) =
   599     by (simp add: Rings.minus_mod_eq_mult_div Groups.mult.commute [of 2] Euclidean_Division.div_mult2_eq \<open>t = 2 * s\<close>)
   683         (2 * (numeral m div (2 * numeral n)) + 1, numeral m mod (2 * numeral n) - numeral n)"
   600       (use mod_mult2_eq [of r s 2] in \<open>simp add: ac_simps \<open>t = 2 * s\<close>\<close>)
   684         by simp
   601   have rs: \<open>r div s mod 2 = 0 \<or> r div s mod 2 = Suc 0\<close>
   685     moreover from True divmod_digit_1 [of "numeral m" "numeral n"]
   602     by auto
   686       have "2 * (numeral m div (2 * numeral n)) + 1 = numeral m div numeral n"
   603   from \<open>\<not> s \<le> r mod s\<close> have \<open>s \<le> r mod t \<Longrightarrow>
   687       and "numeral m mod (2 * numeral n) - numeral n = numeral m mod numeral n"
   604      r div s = Suc (2 * (r div t)) \<and>
   688       by simp_all
   605      r mod s = r mod t - s\<close>
   689     ultimately show ?thesis by (simp only: divmod_def)
   606     using rs
   690   next
   607     by (auto simp add: t)
   691     case False then have *: "numeral m mod (2 * numeral n) < numeral n"
   608   moreover have \<open>r mod t < s \<Longrightarrow>
   692       by (simp add: not_le)
   609      r div s = 2 * (r div t) \<and>
   693     with divmod_step_eq
   610      r mod s = r mod t\<close>
   694       have "divmod_step (numeral n) (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) =
   611     using rs
   695         (2 * (numeral m div (2 * numeral n)), numeral m mod (2 * numeral n))"
   612     by (auto simp add: t)
   696         by auto
   613   ultimately show ?thesis
   697     moreover from * divmod_digit_0 [of "numeral n" "numeral m"]
   614     by (simp add: divmod_def prod_eq_iff split_def Let_def
   698       have "2 * (numeral m div (2 * numeral n)) = numeral m div numeral n"
   615 	    not_less mod_eq_0_iff_dvd Rings.mod_eq_0_iff_dvd False not_le *)
   699       and "numeral m mod (2 * numeral n) = numeral m mod numeral n"
   616     (simp add: flip: of_nat_numeral of_nat_mult add.commute [of 1] of_nat_div of_nat_mod of_nat_Suc of_nat_diff)
   700       by (simp_all only: zero_less_numeral)
       
   701     ultimately show ?thesis by (simp only: divmod_def)
       
   702   qed
       
   703   then have "divmod m n =
       
   704     divmod_step (numeral n) (numeral m div numeral (Num.Bit0 n),
       
   705       numeral m mod numeral (Num.Bit0 n))"
       
   706     by (simp only: numeral.simps distrib mult_1)
       
   707   then have "divmod m n = divmod_step (numeral n) (divmod m (Num.Bit0 n))"
       
   708     by (simp add: divmod_def)
       
   709   with False show ?thesis by simp
       
   710 qed
   617 qed
   711 
   618 
   712 text \<open>The division rewrite proper -- first, trivial results involving \<open>1\<close>\<close>
   619 text \<open>The division rewrite proper -- first, trivial results involving \<open>1\<close>\<close>
   713 
   620 
   714 lemma divmod_trivial [simp]:
   621 lemma divmod_trivial [simp]:
   718   using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def)
   625   using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def)
   719 
   626 
   720 text \<open>Division by an even number is a right-shift\<close>
   627 text \<open>Division by an even number is a right-shift\<close>
   721 
   628 
   722 lemma divmod_cancel [simp]:
   629 lemma divmod_cancel [simp]:
   723   "divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r))" (is ?P)
   630   \<open>divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r))\<close> (is ?P)
   724   "divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r + 1))" (is ?Q)
   631   \<open>divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r + 1))\<close> (is ?Q)
   725 proof -
   632 proof -
   726   have *: "\<And>q. numeral (Num.Bit0 q) = 2 * numeral q"
   633   define r s where \<open>r = (numeral m :: nat)\<close> \<open>s = (numeral n :: nat)\<close>
   727     "\<And>q. numeral (Num.Bit1 q) = 2 * numeral q + 1"
   634   then have *: \<open>numeral m = of_nat r\<close> \<open>numeral n = of_nat s\<close>
   728     by (simp_all only: numeral_mult numeral.simps distrib) simp_all
   635     \<open>numeral (num.Bit0 m) = of_nat (2 * r)\<close> \<open>numeral (num.Bit0 n) = of_nat (2 * s)\<close>
   729   have "1 div 2 = 0" "1 mod 2 = 1" by (auto intro: div_less mod_less)
   636     \<open>numeral (num.Bit1 m) = of_nat (Suc (2 * r))\<close>
   730   then show ?P and ?Q
   637     by simp_all
   731     by (simp_all add: fst_divmod snd_divmod prod_eq_iff split_def * [of m] * [of n] mod_mult_mult1
   638   show ?P and ?Q
   732       div_mult2_eq [of _ _ 2] mod_mult2_eq [of _ _ 2]
   639     by (simp_all add: divmod_def *)
   733       add.commute del: numeral_times_numeral)
   640       (simp_all flip: of_nat_numeral of_nat_div of_nat_mod of_nat_mult add.commute [of 1] of_nat_Suc
       
   641        add: Euclidean_Division.mod_mult_mult1 div_mult2_eq [of _ 2] mod_mult2_eq [of _ 2])
   734 qed
   642 qed
   735 
   643 
   736 text \<open>The really hard work\<close>
   644 text \<open>The really hard work\<close>
   737 
   645 
   738 lemma divmod_steps [simp]:
   646 lemma divmod_steps [simp]:
   746        else divmod_step (numeral (num.Bit1 n))
   654        else divmod_step (numeral (num.Bit1 n))
   747              (divmod (num.Bit1 m)
   655              (divmod (num.Bit1 m)
   748                (num.Bit0 (num.Bit1 n))))"
   656                (num.Bit0 (num.Bit1 n))))"
   749   by (simp_all add: divmod_divmod_step)
   657   by (simp_all add: divmod_divmod_step)
   750 
   658 
   751 lemmas divmod_algorithm_code = divmod_step_eq divmod_trivial divmod_cancel divmod_steps  
   659 lemmas divmod_algorithm_code = divmod_trivial divmod_cancel divmod_steps
   752 
   660 
   753 text \<open>Special case: divisibility\<close>
   661 text \<open>Special case: divisibility\<close>
   754 
   662 
   755 definition divides_aux :: "'a \<times> 'a \<Rightarrow> bool"
   663 definition divides_aux :: "'a \<times> 'a \<Rightarrow> bool"
   756 where
   664 where
   813     \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)"
   721     \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)"
   814   by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even])
   722   by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even])
   815 
   723 
   816 end
   724 end
   817 
   725 
   818 hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq
   726 instantiation nat :: unique_euclidean_semiring_with_nat_division
   819 
       
   820 instantiation nat :: unique_euclidean_semiring_numeral
       
   821 begin
   727 begin
   822 
   728 
   823 definition divmod_nat :: "num \<Rightarrow> num \<Rightarrow> nat \<times> nat"
   729 definition divmod_nat :: "num \<Rightarrow> num \<Rightarrow> nat \<times> nat"
   824 where
   730 where
   825   divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)"
   731   divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)"
   828 where
   734 where
   829   "divmod_step_nat l qr = (let (q, r) = qr
   735   "divmod_step_nat l qr = (let (q, r) = qr
   830     in if r \<ge> l then (2 * q + 1, r - l)
   736     in if r \<ge> l then (2 * q + 1, r - l)
   831     else (2 * q, r))"
   737     else (2 * q, r))"
   832 
   738 
   833 instance by standard
   739 instance
   834   (auto simp add: divmod'_nat_def divmod_step_nat_def div_greater_zero_iff div_mult2_eq mod_mult2_eq)
   740   by standard (simp_all add: divmod'_nat_def divmod_step_nat_def)
   835 
   741 
   836 end
   742 end
   837 
   743 
   838 declare divmod_algorithm_code [where ?'a = nat, code]
   744 declare divmod_algorithm_code [where ?'a = nat, code]
   839 
   745 
   845 lemma Suc_0_mod_numeral [simp]:
   751 lemma Suc_0_mod_numeral [simp]:
   846   fixes k l :: num
   752   fixes k l :: num
   847   shows "Suc 0 mod numeral k = snd (divmod Num.One k)"
   753   shows "Suc 0 mod numeral k = snd (divmod Num.One k)"
   848   by (simp_all add: snd_divmod)
   754   by (simp_all add: snd_divmod)
   849 
   755 
   850 instantiation int :: unique_euclidean_semiring_numeral
   756 instantiation int :: unique_euclidean_semiring_with_nat_division
   851 begin
   757 begin
   852 
   758 
   853 definition divmod_int :: "num \<Rightarrow> num \<Rightarrow> int \<times> int"
   759 definition divmod_int :: "num \<Rightarrow> num \<Rightarrow> int \<times> int"
   854 where
   760 where
   855   "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)"
   761   "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)"
   856 
   762 
   857 definition divmod_step_int :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int"
   763 definition divmod_step_int :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int"
   858 where
   764 where
   859   "divmod_step_int l qr = (let (q, r) = qr
   765   "divmod_step_int l qr = (let (q, r) = qr
   860     in if r \<ge> l then (2 * q + 1, r - l)
   766     in if \<bar>l\<bar> \<le> \<bar>r\<bar> then (2 * q + 1, r - l)
   861     else (2 * q, r))"
   767     else (2 * q, r))"
   862 
   768 
   863 instance
   769 instance
   864   by standard (auto intro: zmod_le_nonneg_dividend simp add: divmod_int_def divmod_step_int_def
   770   by standard (auto simp add: divmod_int_def divmod_step_int_def)
   865     pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq)
       
   866 
   771 
   867 end
   772 end
   868 
   773 
   869 declare divmod_algorithm_code [where ?'a = int, code]
   774 declare divmod_algorithm_code [where ?'a = int, code]
   870 
   775 
   953   \<open>divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\<close>
   858   \<open>divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\<close>
   954   by (cases m) simp_all
   859   by (cases m) simp_all
   955 
   860 
   956 lemma div_positive_int:
   861 lemma div_positive_int:
   957   "k div l > 0" if "k \<ge> l" and "l > 0" for k l :: int
   862   "k div l > 0" if "k \<ge> l" and "l > 0" for k l :: int
   958   using that div_positive [of l k] by blast
   863   using that by (simp add: nonneg1_imp_zdiv_pos_iff)
   959 
   864 
   960 
   865 
   961 subsubsection \<open>Dedicated simproc for calculation\<close>
   866 subsubsection \<open>Dedicated simproc for calculation\<close>
       
   867 
       
   868 lemma euclidean_size_nat_less_eq_iff:
       
   869   \<open>euclidean_size m \<le> euclidean_size n \<longleftrightarrow> m \<le> n\<close> for m n :: nat
       
   870   by simp
       
   871 
       
   872 lemma euclidean_size_int_less_eq_iff:
       
   873   \<open>euclidean_size k \<le> euclidean_size l \<longleftrightarrow> \<bar>k\<bar> \<le> \<bar>l\<bar>\<close> for k l :: int
       
   874   by auto
   962 
   875 
   963 text \<open>
   876 text \<open>
   964   There is space for improvement here: the calculation itself
   877   There is space for improvement here: the calculation itself
   965   could be carried out outside the logic, and a generic simproc
   878   could be carried out outside the logic, and a generic simproc
   966   (simplifier setup) for generic calculation would be helpful. 
   879   (simplifier setup) for generic calculation would be helpful. 
   967 \<close>
   880 \<close>
   968 
   881 
   969 simproc_setup numeral_divmod
   882 simproc_setup numeral_divmod
   970   ("0 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 0 :: 'a :: unique_euclidean_semiring_numeral" |
   883   ("0 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" |
   971    "0 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 1 :: 'a :: unique_euclidean_semiring_numeral" |
   884    "0 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" |
   972    "0 div - 1 :: int" | "0 mod - 1 :: int" |
   885    "0 div - 1 :: int" | "0 mod - 1 :: int" |
   973    "0 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" |
   886    "0 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" |
   974    "0 div - numeral b :: int" | "0 mod - numeral b :: int" |
   887    "0 div - numeral b :: int" | "0 mod - numeral b :: int" |
   975    "1 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 0 :: 'a :: unique_euclidean_semiring_numeral" |
   888    "1 div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" |
   976    "1 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 1 :: 'a :: unique_euclidean_semiring_numeral" |
   889    "1 div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" |
   977    "1 div - 1 :: int" | "1 mod - 1 :: int" |
   890    "1 div - 1 :: int" | "1 mod - 1 :: int" |
   978    "1 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" |
   891    "1 div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" |
   979    "1 div - numeral b :: int" |"1 mod - numeral b :: int" |
   892    "1 div - numeral b :: int" |"1 mod - numeral b :: int" |
   980    "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" |
   893    "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" |
   981    "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" |
   894    "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" |
   982    "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" |
   895    "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" |
   983    "numeral a div 0 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_numeral" |
   896    "numeral a div 0 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_with_nat_division" |
   984    "numeral a div 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_numeral" |
   897    "numeral a div 1 :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_with_nat_division" |
   985    "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" |
   898    "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" |
   986    "numeral a div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_numeral" |
   899    "numeral a div numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_with_nat_division" |
   987    "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" |
   900    "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" |
   988    "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" |
   901    "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" |
   989    "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" |
   902    "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" |
   990    "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" |
   903    "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" |
   991    "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" |
   904    "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" |
  1003         one_div_minus_numeral one_mod_minus_numeral
   916         one_div_minus_numeral one_mod_minus_numeral
  1004         numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral
   917         numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral
  1005         numeral_div_minus_numeral numeral_mod_minus_numeral
   918         numeral_div_minus_numeral numeral_mod_minus_numeral
  1006         div_minus_minus mod_minus_minus Divides.adjust_div_eq of_bool_eq one_neq_zero
   919         div_minus_minus mod_minus_minus Divides.adjust_div_eq of_bool_eq one_neq_zero
  1007         numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial
   920         numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial
  1008         divmod_cancel divmod_steps divmod_step_eq fst_conv snd_conv numeral_One
   921         divmod_cancel divmod_steps divmod_step_def fst_conv snd_conv numeral_One
  1009         case_prod_beta rel_simps Divides.adjust_mod_def div_minus1_right mod_minus1_right
   922         case_prod_beta rel_simps Divides.adjust_mod_def div_minus1_right mod_minus1_right
  1010         minus_minus numeral_times_numeral mult_zero_right mult_1_right}
   923         minus_minus numeral_times_numeral mult_zero_right mult_1_right
       
   924         euclidean_size_nat_less_eq_iff euclidean_size_int_less_eq_iff diff_nat_numeral nat_numeral}
  1011         @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}]);
   925         @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}]);
  1012       fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt
   926       fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt
  1013         (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps)
   927         (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps)
  1014     in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end
   928     in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end
  1015   end
   929   end
  1064   code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
   978   code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
  1065 
   979 
  1066 
   980 
  1067 subsection \<open>Lemmas of doubtful value\<close>
   981 subsection \<open>Lemmas of doubtful value\<close>
  1068 
   982 
       
   983 class unique_euclidean_semiring_numeral = unique_euclidean_semiring_with_nat + linordered_semidom +
       
   984   assumes div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0"
       
   985     and mod_less: " 0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a mod b = a"
       
   986     and div_positive: "0 < b \<Longrightarrow> b \<le> a \<Longrightarrow> a div b > 0"
       
   987     and mod_less_eq_dividend: "0 \<le> a \<Longrightarrow> a mod b \<le> a"
       
   988     and pos_mod_bound: "0 < b \<Longrightarrow> a mod b < b"
       
   989     and pos_mod_sign: "0 < b \<Longrightarrow> 0 \<le> a mod b"
       
   990     and mod_mult2_eq: "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b"
       
   991     and div_mult2_eq: "0 \<le> c \<Longrightarrow> a div (b * c) = a div b div c"
       
   992   assumes discrete: "a < b \<longleftrightarrow> a + 1 \<le> b"
       
   993 begin
       
   994 
       
   995 lemma divmod_digit_1:
       
   996   assumes "0 \<le> a" "0 < b" and "b \<le> a mod (2 * b)"
       
   997   shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P")
       
   998     and "a mod (2 * b) - b = a mod b" (is "?Q")
       
   999 proof -
       
  1000   from assms mod_less_eq_dividend [of a "2 * b"] have "b \<le> a"
       
  1001     by (auto intro: trans)
       
  1002   with \<open>0 < b\<close> have "0 < a div b" by (auto intro: div_positive)
       
  1003   then have [simp]: "1 \<le> a div b" by (simp add: discrete)
       
  1004   with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound)
       
  1005   define w where "w = a div b mod 2"
       
  1006   then have w_exhaust: "w = 0 \<or> w = 1" by auto
       
  1007   have mod_w: "a mod (2 * b) = a mod b + b * w"
       
  1008     by (simp add: w_def mod_mult2_eq ac_simps)
       
  1009   from assms w_exhaust have "w = 1"
       
  1010     using mod_less by (auto simp add: mod_w)
       
  1011   with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp
       
  1012   have "2 * (a div (2 * b)) = a div b - w"
       
  1013     by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps)
       
  1014   with \<open>w = 1\<close> have div: "2 * (a div (2 * b)) = a div b - 1" by simp
       
  1015   then show ?P and ?Q
       
  1016     by (simp_all add: div mod add_implies_diff [symmetric])
       
  1017 qed
       
  1018 
       
  1019 lemma divmod_digit_0:
       
  1020   assumes "0 < b" and "a mod (2 * b) < b"
       
  1021   shows "2 * (a div (2 * b)) = a div b" (is "?P")
       
  1022     and "a mod (2 * b) = a mod b" (is "?Q")
       
  1023 proof -
       
  1024   define w where "w = a div b mod 2"
       
  1025   then have w_exhaust: "w = 0 \<or> w = 1" by auto
       
  1026   have mod_w: "a mod (2 * b) = a mod b + b * w"
       
  1027     by (simp add: w_def mod_mult2_eq ac_simps)
       
  1028   moreover have "b \<le> a mod b + b"
       
  1029   proof -
       
  1030     from \<open>0 < b\<close> pos_mod_sign have "0 \<le> a mod b" by blast
       
  1031     then have "0 + b \<le> a mod b + b" by (rule add_right_mono)
       
  1032     then show ?thesis by simp
       
  1033   qed
       
  1034   moreover note assms w_exhaust
       
  1035   ultimately have "w = 0" by auto
       
  1036   with mod_w have mod: "a mod (2 * b) = a mod b" by simp
       
  1037   have "2 * (a div (2 * b)) = a div b - w"
       
  1038     by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps)
       
  1039   with \<open>w = 0\<close> have div: "2 * (a div (2 * b)) = a div b" by simp
       
  1040   then show ?P and ?Q
       
  1041     by (simp_all add: div mod)
       
  1042 qed
       
  1043 
       
  1044 lemma mod_double_modulus:
       
  1045   assumes "m > 0" "x \<ge> 0"
       
  1046   shows   "x mod (2 * m) = x mod m \<or> x mod (2 * m) = x mod m + m"
       
  1047 proof (cases "x mod (2 * m) < m")
       
  1048   case True
       
  1049   thus ?thesis using assms using divmod_digit_0(2)[of m x] by auto
       
  1050 next
       
  1051   case False
       
  1052   hence *: "x mod (2 * m) - m = x mod m"
       
  1053     using assms by (intro divmod_digit_1) auto
       
  1054   hence "x mod (2 * m) = x mod m + m"
       
  1055     by (subst * [symmetric], subst le_add_diff_inverse2) (use False in auto)
       
  1056   thus ?thesis by simp
       
  1057 qed
       
  1058 
       
  1059 end
       
  1060 
       
  1061 hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq
       
  1062 
       
  1063 instance nat :: unique_euclidean_semiring_numeral
       
  1064   by standard
       
  1065     (auto simp add: div_greater_zero_iff div_mult2_eq mod_mult2_eq)
       
  1066 
       
  1067 instance int :: unique_euclidean_semiring_numeral
       
  1068   by standard (auto intro: zmod_le_nonneg_dividend simp add:
       
  1069     pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq)
       
  1070 
  1069 lemma div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and " \<not> m < n" for m n :: nat
  1071 lemma div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and " \<not> m < n" for m n :: nat
  1070   by (rule le_div_geq) (use that in \<open>simp_all add: not_less\<close>)
  1072   by (rule le_div_geq) (use that in \<open>simp_all add: not_less\<close>)
  1071 
  1073 
  1072 lemma mod_geq: "m mod n = (m - n) mod n" if "\<not> m < n" for m n :: nat
  1074 lemma mod_geq: "m mod n = (m - n) mod n" if "\<not> m < n" for m n :: nat
  1073   by (rule le_mod_geq) (use that in \<open>simp add: not_less\<close>)
  1075   by (rule le_mod_geq) (use that in \<open>simp add: not_less\<close>)