| author | bulwahn | 
| Sat, 20 Oct 2012 09:09:32 +0200 | |
| changeset 49943 | 6a26fed71755 | 
| parent 48891 | c0eafbd55de3 | 
| child 49962 | a8cc904a6820 | 
| permissions | -rw-r--r-- | 
| 3366 | 1 | (* Title: HOL/Divides.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 3 | Copyright 1999 University of Cambridge | 
| 18154 | 4 | *) | 
| 3366 | 5 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 6 | header {* The division operators div and mod *}
 | 
| 3366 | 7 | |
| 15131 | 8 | theory Divides | 
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47217diff
changeset | 9 | imports Nat_Transfer | 
| 15131 | 10 | begin | 
| 3366 | 11 | |
| 48891 | 12 | ML_file "~~/src/Provers/Arith/cancel_div_mod.ML" | 
| 13 | ||
| 25942 | 14 | subsection {* Syntactic division operations *}
 | 
| 15 | ||
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 16 | class div = dvd + | 
| 27540 | 17 | fixes div :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "div" 70) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 18 | and mod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "mod" 70) | 
| 27540 | 19 | |
| 20 | ||
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 21 | subsection {* Abstract division in commutative semirings. *}
 | 
| 25942 | 22 | |
| 30930 | 23 | class semiring_div = comm_semiring_1_cancel + no_zero_divisors + div + | 
| 25942 | 24 | assumes mod_div_equality: "a div b * b + a mod b = a" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 25 | and div_by_0 [simp]: "a div 0 = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 26 | and div_0 [simp]: "0 div a = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 27 | and div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b" | 
| 30930 | 28 | and div_mult_mult1 [simp]: "c \<noteq> 0 \<Longrightarrow> (c * a) div (c * b) = a div b" | 
| 25942 | 29 | begin | 
| 30 | ||
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 31 | text {* @{const div} and @{const mod} *}
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 32 | |
| 26062 | 33 | lemma mod_div_equality2: "b * (a div b) + a mod b = a" | 
| 34 | unfolding mult_commute [of b] | |
| 35 | by (rule mod_div_equality) | |
| 36 | ||
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 37 | lemma mod_div_equality': "a mod b + a div b * b = a" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 38 | using mod_div_equality [of a b] | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 39 | by (simp only: add_ac) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 40 | |
| 26062 | 41 | lemma div_mod_equality: "((a div b) * b + a mod b) + c = a + c" | 
| 30934 | 42 | by (simp add: mod_div_equality) | 
| 26062 | 43 | |
| 44 | lemma div_mod_equality2: "(b * (a div b) + a mod b) + c = a + c" | |
| 30934 | 45 | by (simp add: mod_div_equality2) | 
| 26062 | 46 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 47 | lemma mod_by_0 [simp]: "a mod 0 = a" | 
| 30934 | 48 | using mod_div_equality [of a zero] by simp | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 49 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 50 | lemma mod_0 [simp]: "0 mod a = 0" | 
| 30934 | 51 | using mod_div_equality [of zero a] div_0 by simp | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 52 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 53 | lemma div_mult_self2 [simp]: | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 54 | assumes "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 55 | shows "(a + b * c) div b = c + a div b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 56 | using assms div_mult_self1 [of b a c] by (simp add: mult_commute) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 57 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 58 | lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 59 | proof (cases "b = 0") | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 60 | case True then show ?thesis by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 61 | next | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 62 | case False | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 63 | have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 64 | by (simp add: mod_div_equality) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 65 | also from False div_mult_self1 [of b a c] have | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 66 | "\<dots> = (c + a div b) * b + (a + c * b) mod b" | 
| 29667 | 67 | by (simp add: algebra_simps) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 68 | finally have "a = a div b * b + (a + c * b) mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 69 | by (simp add: add_commute [of a] add_assoc left_distrib) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 70 | then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 71 | by (simp add: mod_div_equality) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 72 | then show ?thesis by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 73 | qed | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 74 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 75 | lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" | 
| 30934 | 76 | by (simp add: mult_commute [of b]) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 77 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 78 | lemma div_mult_self1_is_id [simp]: "b \<noteq> 0 \<Longrightarrow> b * a div b = a" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 79 | using div_mult_self2 [of b 0 a] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 80 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 81 | lemma div_mult_self2_is_id [simp]: "b \<noteq> 0 \<Longrightarrow> a * b div b = a" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 82 | using div_mult_self1 [of b 0 a] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 83 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 84 | lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 85 | using mod_mult_self2 [of 0 b a] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 86 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 87 | lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 88 | using mod_mult_self1 [of 0 a b] by simp | 
| 26062 | 89 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 90 | lemma div_by_1 [simp]: "a div 1 = a" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 91 | using div_mult_self2_is_id [of 1 a] zero_neq_one by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 92 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 93 | lemma mod_by_1 [simp]: "a mod 1 = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 94 | proof - | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 95 | from mod_div_equality [of a one] div_by_1 have "a + a mod 1 = a" by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 96 | then have "a + a mod 1 = a + 0" by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 97 | then show ?thesis by (rule add_left_imp_eq) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 98 | qed | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 99 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 100 | lemma mod_self [simp]: "a mod a = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 101 | using mod_mult_self2_is_0 [of 1] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 102 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 103 | lemma div_self [simp]: "a \<noteq> 0 \<Longrightarrow> a div a = 1" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 104 | using div_mult_self2_is_id [of _ 1] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 105 | |
| 27676 | 106 | lemma div_add_self1 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 107 | assumes "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 108 | shows "(b + a) div b = a div b + 1" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 109 | using assms div_mult_self1 [of b a 1] by (simp add: add_commute) | 
| 26062 | 110 | |
| 27676 | 111 | lemma div_add_self2 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 112 | assumes "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 113 | shows "(a + b) div b = a div b + 1" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 114 | using assms div_add_self1 [of b a] by (simp add: add_commute) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 115 | |
| 27676 | 116 | lemma mod_add_self1 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 117 | "(b + a) mod b = a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 118 | using mod_mult_self1 [of a 1 b] by (simp add: add_commute) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 119 | |
| 27676 | 120 | lemma mod_add_self2 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 121 | "(a + b) mod b = a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 122 | using mod_mult_self1 [of a 1 b] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 123 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 124 | lemma mod_div_decomp: | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 125 | fixes a b | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 126 | obtains q r where "q = a div b" and "r = a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 127 | and "a = q * b + r" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 128 | proof - | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 129 | from mod_div_equality have "a = a div b * b + a mod b" by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 130 | moreover have "a div b = a div b" .. | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 131 | moreover have "a mod b = a mod b" .. | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 132 | note that ultimately show thesis by blast | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 133 | qed | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 134 | |
| 45231 
d85a2fdc586c
replacing code_inline by code_unfold, removing obsolete code_unfold, code_inline del now that the ancient code generator is removed
 bulwahn parents: 
44890diff
changeset | 135 | lemma dvd_eq_mod_eq_0 [code]: "a dvd b \<longleftrightarrow> b mod a = 0" | 
| 25942 | 136 | proof | 
| 137 | assume "b mod a = 0" | |
| 138 | with mod_div_equality [of b a] have "b div a * a = b" by simp | |
| 139 | then have "b = a * (b div a)" unfolding mult_commute .. | |
| 140 | then have "\<exists>c. b = a * c" .. | |
| 141 | then show "a dvd b" unfolding dvd_def . | |
| 142 | next | |
| 143 | assume "a dvd b" | |
| 144 | then have "\<exists>c. b = a * c" unfolding dvd_def . | |
| 145 | then obtain c where "b = a * c" .. | |
| 146 | then have "b mod a = a * c mod a" by simp | |
| 147 | then have "b mod a = c * a mod a" by (simp add: mult_commute) | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 148 | then show "b mod a = 0" by simp | 
| 25942 | 149 | qed | 
| 150 | ||
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 151 | lemma mod_div_trivial [simp]: "a mod b div b = 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 152 | proof (cases "b = 0") | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 153 | assume "b = 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 154 | thus ?thesis by simp | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 155 | next | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 156 | assume "b \<noteq> 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 157 | hence "a div b + a mod b div b = (a mod b + a div b * b) div b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 158 | by (rule div_mult_self1 [symmetric]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 159 | also have "\<dots> = a div b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 160 | by (simp only: mod_div_equality') | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 161 | also have "\<dots> = a div b + 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 162 | by simp | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 163 | finally show ?thesis | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 164 | by (rule add_left_imp_eq) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 165 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 166 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 167 | lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 168 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 169 | have "a mod b mod b = (a mod b + a div b * b) mod b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 170 | by (simp only: mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 171 | also have "\<dots> = a mod b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 172 | by (simp only: mod_div_equality') | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 173 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 174 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 175 | |
| 29925 | 176 | lemma dvd_imp_mod_0: "a dvd b \<Longrightarrow> b mod a = 0" | 
| 29948 | 177 | by (rule dvd_eq_mod_eq_0[THEN iffD1]) | 
| 29925 | 178 | |
| 179 | lemma dvd_div_mult_self: "a dvd b \<Longrightarrow> (b div a) * a = b" | |
| 180 | by (subst (2) mod_div_equality [of b a, symmetric]) (simp add:dvd_imp_mod_0) | |
| 181 | ||
| 33274 
b6ff7db522b5
moved lemmas for dvd on nat to theories Nat and Power
 haftmann parents: 
32010diff
changeset | 182 | lemma dvd_mult_div_cancel: "a dvd b \<Longrightarrow> a * (b div a) = b" | 
| 
b6ff7db522b5
moved lemmas for dvd on nat to theories Nat and Power
 haftmann parents: 
32010diff
changeset | 183 | by (drule dvd_div_mult_self) (simp add: mult_commute) | 
| 
b6ff7db522b5
moved lemmas for dvd on nat to theories Nat and Power
 haftmann parents: 
32010diff
changeset | 184 | |
| 30052 | 185 | lemma dvd_div_mult: "a dvd b \<Longrightarrow> (b div a) * c = b * c div a" | 
| 186 | apply (cases "a = 0") | |
| 187 | apply simp | |
| 188 | apply (auto simp: dvd_def mult_assoc) | |
| 189 | done | |
| 190 | ||
| 29925 | 191 | lemma div_dvd_div[simp]: | 
| 192 | "a dvd b \<Longrightarrow> a dvd c \<Longrightarrow> (b div a dvd c div a) = (b dvd c)" | |
| 193 | apply (cases "a = 0") | |
| 194 | apply simp | |
| 195 | apply (unfold dvd_def) | |
| 196 | apply auto | |
| 197 | apply(blast intro:mult_assoc[symmetric]) | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44766diff
changeset | 198 | apply(fastforce simp add: mult_assoc) | 
| 29925 | 199 | done | 
| 200 | ||
| 30078 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 201 | lemma dvd_mod_imp_dvd: "[| k dvd m mod n; k dvd n |] ==> k dvd m" | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 202 | apply (subgoal_tac "k dvd (m div n) *n + m mod n") | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 203 | apply (simp add: mod_div_equality) | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 204 | apply (simp only: dvd_add dvd_mult) | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 205 | done | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 206 | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 207 | text {* Addition respects modular equivalence. *}
 | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 208 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 209 | lemma mod_add_left_eq: "(a + b) mod c = (a mod c + b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 210 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 211 | have "(a + b) mod c = (a div c * c + a mod c + b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 212 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 213 | also have "\<dots> = (a mod c + b + a div c * c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 214 | by (simp only: add_ac) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 215 | also have "\<dots> = (a mod c + b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 216 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 217 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 218 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 219 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 220 | lemma mod_add_right_eq: "(a + b) mod c = (a + b mod c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 221 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 222 | have "(a + b) mod c = (a + (b div c * c + b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 223 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 224 | also have "\<dots> = (a + b mod c + b div c * c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 225 | by (simp only: add_ac) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 226 | also have "\<dots> = (a + b mod c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 227 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 228 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 229 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 230 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 231 | lemma mod_add_eq: "(a + b) mod c = (a mod c + b mod c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 232 | by (rule trans [OF mod_add_left_eq mod_add_right_eq]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 233 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 234 | lemma mod_add_cong: | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 235 | assumes "a mod c = a' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 236 | assumes "b mod c = b' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 237 | shows "(a + b) mod c = (a' + b') mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 238 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 239 | have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 240 | unfolding assms .. | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 241 | thus ?thesis | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 242 | by (simp only: mod_add_eq [symmetric]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 243 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 244 | |
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 245 | lemma div_add [simp]: "z dvd x \<Longrightarrow> z dvd y | 
| 30837 
3d4832d9f7e4
added strong_setprod_cong[cong] (in analogy with setsum)
 nipkow parents: 
30729diff
changeset | 246 | \<Longrightarrow> (x + y) div z = x div z + y div z" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 247 | by (cases "z = 0", simp, unfold dvd_def, auto simp add: algebra_simps) | 
| 30837 
3d4832d9f7e4
added strong_setprod_cong[cong] (in analogy with setsum)
 nipkow parents: 
30729diff
changeset | 248 | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 249 | text {* Multiplication respects modular equivalence. *}
 | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 250 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 251 | lemma mod_mult_left_eq: "(a * b) mod c = ((a mod c) * b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 252 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 253 | have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 254 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 255 | also have "\<dots> = (a mod c * b + a div c * b * c) mod c" | 
| 29667 | 256 | by (simp only: algebra_simps) | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 257 | also have "\<dots> = (a mod c * b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 258 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 259 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 260 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 261 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 262 | lemma mod_mult_right_eq: "(a * b) mod c = (a * (b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 263 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 264 | have "(a * b) mod c = (a * (b div c * c + b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 265 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 266 | also have "\<dots> = (a * (b mod c) + a * (b div c) * c) mod c" | 
| 29667 | 267 | by (simp only: algebra_simps) | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 268 | also have "\<dots> = (a * (b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 269 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 270 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 271 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 272 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 273 | lemma mod_mult_eq: "(a * b) mod c = ((a mod c) * (b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 274 | by (rule trans [OF mod_mult_left_eq mod_mult_right_eq]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 275 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 276 | lemma mod_mult_cong: | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 277 | assumes "a mod c = a' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 278 | assumes "b mod c = b' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 279 | shows "(a * b) mod c = (a' * b') mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 280 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 281 | have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 282 | unfolding assms .. | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 283 | thus ?thesis | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 284 | by (simp only: mod_mult_eq [symmetric]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 285 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 286 | |
| 47164 | 287 | text {* Exponentiation respects modular equivalence. *}
 | 
| 288 | ||
| 289 | lemma power_mod: "(a mod b)^n mod b = a^n mod b" | |
| 290 | apply (induct n, simp_all) | |
| 291 | apply (rule mod_mult_right_eq [THEN trans]) | |
| 292 | apply (simp (no_asm_simp)) | |
| 293 | apply (rule mod_mult_eq [symmetric]) | |
| 294 | done | |
| 295 | ||
| 29404 | 296 | lemma mod_mod_cancel: | 
| 297 | assumes "c dvd b" | |
| 298 | shows "a mod b mod c = a mod c" | |
| 299 | proof - | |
| 300 | from `c dvd b` obtain k where "b = c * k" | |
| 301 | by (rule dvdE) | |
| 302 | have "a mod b mod c = a mod (c * k) mod c" | |
| 303 | by (simp only: `b = c * k`) | |
| 304 | also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c" | |
| 305 | by (simp only: mod_mult_self1) | |
| 306 | also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c" | |
| 307 | by (simp only: add_ac mult_ac) | |
| 308 | also have "\<dots> = a mod c" | |
| 309 | by (simp only: mod_div_equality) | |
| 310 | finally show ?thesis . | |
| 311 | qed | |
| 312 | ||
| 30930 | 313 | lemma div_mult_div_if_dvd: | 
| 314 | "y dvd x \<Longrightarrow> z dvd w \<Longrightarrow> (x div y) * (w div z) = (x * w) div (y * z)" | |
| 315 | apply (cases "y = 0", simp) | |
| 316 | apply (cases "z = 0", simp) | |
| 317 | apply (auto elim!: dvdE simp add: algebra_simps) | |
| 30476 | 318 | apply (subst mult_assoc [symmetric]) | 
| 319 | apply (simp add: no_zero_divisors) | |
| 30930 | 320 | done | 
| 321 | ||
| 35367 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 322 | lemma div_mult_swap: | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 323 | assumes "c dvd b" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 324 | shows "a * (b div c) = (a * b) div c" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 325 | proof - | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 326 | from assms have "b div c * (a div 1) = b * a div (c * 1)" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 327 | by (simp only: div_mult_div_if_dvd one_dvd) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 328 | then show ?thesis by (simp add: mult_commute) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 329 | qed | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 330 | |
| 30930 | 331 | lemma div_mult_mult2 [simp]: | 
| 332 | "c \<noteq> 0 \<Longrightarrow> (a * c) div (b * c) = a div b" | |
| 333 | by (drule div_mult_mult1) (simp add: mult_commute) | |
| 334 | ||
| 335 | lemma div_mult_mult1_if [simp]: | |
| 336 | "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" | |
| 337 | by simp_all | |
| 30476 | 338 | |
| 30930 | 339 | lemma mod_mult_mult1: | 
| 340 | "(c * a) mod (c * b) = c * (a mod b)" | |
| 341 | proof (cases "c = 0") | |
| 342 | case True then show ?thesis by simp | |
| 343 | next | |
| 344 | case False | |
| 345 | from mod_div_equality | |
| 346 | have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . | |
| 347 | with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) | |
| 348 | = c * a + c * (a mod b)" by (simp add: algebra_simps) | |
| 349 | with mod_div_equality show ?thesis by simp | |
| 350 | qed | |
| 351 | ||
| 352 | lemma mod_mult_mult2: | |
| 353 | "(a * c) mod (b * c) = (a mod b) * c" | |
| 354 | using mod_mult_mult1 [of c a b] by (simp add: mult_commute) | |
| 355 | ||
| 47159 | 356 | lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" | 
| 357 | by (fact mod_mult_mult2 [symmetric]) | |
| 358 | ||
| 359 | lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)" | |
| 360 | by (fact mod_mult_mult1 [symmetric]) | |
| 361 | ||
| 31662 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 362 | lemma dvd_mod: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m mod n)" | 
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 363 | unfolding dvd_def by (auto simp add: mod_mult_mult1) | 
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 364 | |
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 365 | lemma dvd_mod_iff: "k dvd n \<Longrightarrow> k dvd (m mod n) \<longleftrightarrow> k dvd m" | 
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 366 | by (blast intro: dvd_mod_imp_dvd dvd_mod) | 
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 367 | |
| 31009 
41fd307cab30
dropped reference to class recpower and lemma duplicate
 haftmann parents: 
30934diff
changeset | 368 | lemma div_power: | 
| 31661 
1e252b8b2334
move lemma div_power into semiring_div context; class ring_div inherits from idom
 huffman parents: 
31009diff
changeset | 369 | "y dvd x \<Longrightarrow> (x div y) ^ n = x ^ n div y ^ n" | 
| 30476 | 370 | apply (induct n) | 
| 371 | apply simp | |
| 372 | apply(simp add: div_mult_div_if_dvd dvd_power_same) | |
| 373 | done | |
| 374 | ||
| 35367 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 375 | lemma dvd_div_eq_mult: | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 376 | assumes "a \<noteq> 0" and "a dvd b" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 377 | shows "b div a = c \<longleftrightarrow> b = c * a" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 378 | proof | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 379 | assume "b = c * a" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 380 | then show "b div a = c" by (simp add: assms) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 381 | next | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 382 | assume "b div a = c" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 383 | then have "b div a * a = c * a" by simp | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 384 | moreover from `a dvd b` have "b div a * a = b" by (simp add: dvd_div_mult_self) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 385 | ultimately show "b = c * a" by simp | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 386 | qed | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 387 | |
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 388 | lemma dvd_div_div_eq_mult: | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 389 | assumes "a \<noteq> 0" "c \<noteq> 0" and "a dvd b" "c dvd d" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 390 | shows "b div a = d div c \<longleftrightarrow> b * c = a * d" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 391 | using assms by (auto simp add: mult_commute [of _ a] dvd_div_mult_self dvd_div_eq_mult div_mult_swap intro: sym) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 392 | |
| 31661 
1e252b8b2334
move lemma div_power into semiring_div context; class ring_div inherits from idom
 huffman parents: 
31009diff
changeset | 393 | end | 
| 
1e252b8b2334
move lemma div_power into semiring_div context; class ring_div inherits from idom
 huffman parents: 
31009diff
changeset | 394 | |
| 35673 | 395 | class ring_div = semiring_div + comm_ring_1 | 
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 396 | begin | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 397 | |
| 36634 | 398 | subclass ring_1_no_zero_divisors .. | 
| 399 | ||
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 400 | text {* Negation respects modular equivalence. *}
 | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 401 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 402 | lemma mod_minus_eq: "(- a) mod b = (- (a mod b)) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 403 | proof - | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 404 | have "(- a) mod b = (- (a div b * b + a mod b)) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 405 | by (simp only: mod_div_equality) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 406 | also have "\<dots> = (- (a mod b) + - (a div b) * b) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 407 | by (simp only: minus_add_distrib minus_mult_left add_ac) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 408 | also have "\<dots> = (- (a mod b)) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 409 | by (rule mod_mult_self1) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 410 | finally show ?thesis . | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 411 | qed | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 412 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 413 | lemma mod_minus_cong: | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 414 | assumes "a mod b = a' mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 415 | shows "(- a) mod b = (- a') mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 416 | proof - | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 417 | have "(- (a mod b)) mod b = (- (a' mod b)) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 418 | unfolding assms .. | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 419 | thus ?thesis | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 420 | by (simp only: mod_minus_eq [symmetric]) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 421 | qed | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 422 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 423 | text {* Subtraction respects modular equivalence. *}
 | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 424 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 425 | lemma mod_diff_left_eq: "(a - b) mod c = (a mod c - b) mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 426 | unfolding diff_minus | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 427 | by (intro mod_add_cong mod_minus_cong) simp_all | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 428 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 429 | lemma mod_diff_right_eq: "(a - b) mod c = (a - b mod c) mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 430 | unfolding diff_minus | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 431 | by (intro mod_add_cong mod_minus_cong) simp_all | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 432 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 433 | lemma mod_diff_eq: "(a - b) mod c = (a mod c - b mod c) mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 434 | unfolding diff_minus | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 435 | by (intro mod_add_cong mod_minus_cong) simp_all | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 436 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 437 | lemma mod_diff_cong: | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 438 | assumes "a mod c = a' mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 439 | assumes "b mod c = b' mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 440 | shows "(a - b) mod c = (a' - b') mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 441 | unfolding diff_minus using assms | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 442 | by (intro mod_add_cong mod_minus_cong) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 443 | |
| 30180 | 444 | lemma dvd_neg_div: "y dvd x \<Longrightarrow> -x div y = - (x div y)" | 
| 445 | apply (case_tac "y = 0") apply simp | |
| 446 | apply (auto simp add: dvd_def) | |
| 447 | apply (subgoal_tac "-(y * k) = y * - k") | |
| 448 | apply (erule ssubst) | |
| 449 | apply (erule div_mult_self1_is_id) | |
| 450 | apply simp | |
| 451 | done | |
| 452 | ||
| 453 | lemma dvd_div_neg: "y dvd x \<Longrightarrow> x div -y = - (x div y)" | |
| 454 | apply (case_tac "y = 0") apply simp | |
| 455 | apply (auto simp add: dvd_def) | |
| 456 | apply (subgoal_tac "y * k = -y * -k") | |
| 457 | apply (erule ssubst) | |
| 458 | apply (rule div_mult_self1_is_id) | |
| 459 | apply simp | |
| 460 | apply simp | |
| 461 | done | |
| 462 | ||
| 47159 | 463 | lemma div_minus_minus [simp]: "(-a) div (-b) = a div b" | 
| 464 | using div_mult_mult1 [of "- 1" a b] | |
| 465 | unfolding neg_equal_0_iff_equal by simp | |
| 466 | ||
| 467 | lemma mod_minus_minus [simp]: "(-a) mod (-b) = - (a mod b)" | |
| 468 | using mod_mult_mult1 [of "- 1" a b] by simp | |
| 469 | ||
| 470 | lemma div_minus_right: "a div (-b) = (-a) div b" | |
| 471 | using div_minus_minus [of "-a" b] by simp | |
| 472 | ||
| 473 | lemma mod_minus_right: "a mod (-b) = - ((-a) mod b)" | |
| 474 | using mod_minus_minus [of "-a" b] by simp | |
| 475 | ||
| 47160 | 476 | lemma div_minus1_right [simp]: "a div (-1) = -a" | 
| 477 | using div_minus_right [of a 1] by simp | |
| 478 | ||
| 479 | lemma mod_minus1_right [simp]: "a mod (-1) = 0" | |
| 480 | using mod_minus_right [of a 1] by simp | |
| 481 | ||
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 482 | end | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 483 | |
| 25942 | 484 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 485 | subsection {* Division on @{typ nat} *}
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 486 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 487 | text {*
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 488 |   We define @{const div} and @{const mod} on @{typ nat} by means
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 489 | of a characteristic relation with two input arguments | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 490 |   @{term "m\<Colon>nat"}, @{term "n\<Colon>nat"} and two output arguments
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 491 |   @{term "q\<Colon>nat"}(uotient) and @{term "r\<Colon>nat"}(emainder).
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 492 | *} | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 493 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 494 | definition divmod_nat_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat \<Rightarrow> bool" where | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 495 | "divmod_nat_rel m n qr \<longleftrightarrow> | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 496 | m = fst qr * n + snd qr \<and> | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 497 | (if n = 0 then fst qr = 0 else if n > 0 then 0 \<le> snd qr \<and> snd qr < n else n < snd qr \<and> snd qr \<le> 0)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 498 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 499 | text {* @{const divmod_nat_rel} is total: *}
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 500 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 501 | lemma divmod_nat_rel_ex: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 502 | obtains q r where "divmod_nat_rel m n (q, r)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 503 | proof (cases "n = 0") | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 504 | case True with that show thesis | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 505 | by (auto simp add: divmod_nat_rel_def) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 506 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 507 | case False | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 508 | have "\<exists>q r. m = q * n + r \<and> r < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 509 | proof (induct m) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 510 | case 0 with `n \<noteq> 0` | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 511 | have "(0\<Colon>nat) = 0 * n + 0 \<and> 0 < n" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 512 | then show ?case by blast | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 513 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 514 | case (Suc m) then obtain q' r' | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 515 | where m: "m = q' * n + r'" and n: "r' < n" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 516 | then show ?case proof (cases "Suc r' < n") | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 517 | case True | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 518 | from m n have "Suc m = q' * n + Suc r'" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 519 | with True show ?thesis by blast | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 520 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 521 | case False then have "n \<le> Suc r'" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 522 | moreover from n have "Suc r' \<le> n" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 523 | ultimately have "n = Suc r'" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 524 | with m have "Suc m = Suc q' * n + 0" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 525 | with `n \<noteq> 0` show ?thesis by blast | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 526 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 527 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 528 | with that show thesis | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 529 | using `n \<noteq> 0` by (auto simp add: divmod_nat_rel_def) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 530 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 531 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 532 | text {* @{const divmod_nat_rel} is injective: *}
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 533 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 534 | lemma divmod_nat_rel_unique: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 535 | assumes "divmod_nat_rel m n qr" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 536 | and "divmod_nat_rel m n qr'" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 537 | shows "qr = qr'" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 538 | proof (cases "n = 0") | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 539 | case True with assms show ?thesis | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 540 | by (cases qr, cases qr') | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 541 | (simp add: divmod_nat_rel_def) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 542 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 543 | case False | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 544 | have aux: "\<And>q r q' r'. q' * n + r' = q * n + r \<Longrightarrow> r < n \<Longrightarrow> q' \<le> (q\<Colon>nat)" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 545 | apply (rule leI) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 546 | apply (subst less_iff_Suc_add) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 547 | apply (auto simp add: add_mult_distrib) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 548 | done | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 549 | from `n \<noteq> 0` assms have "fst qr = fst qr'" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 550 | by (auto simp add: divmod_nat_rel_def intro: order_antisym dest: aux sym) | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 551 | moreover from this assms have "snd qr = snd qr'" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 552 | by (simp add: divmod_nat_rel_def) | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 553 | ultimately show ?thesis by (cases qr, cases qr') simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 554 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 555 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 556 | text {*
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 557 | We instantiate divisibility on the natural numbers by | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 558 |   means of @{const divmod_nat_rel}:
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 559 | *} | 
| 25942 | 560 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 561 | definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" where | 
| 37767 | 562 | "divmod_nat m n = (THE qr. divmod_nat_rel m n qr)" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 563 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 564 | lemma divmod_nat_rel_divmod_nat: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 565 | "divmod_nat_rel m n (divmod_nat m n)" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 566 | proof - | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 567 | from divmod_nat_rel_ex | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 568 | obtain qr where rel: "divmod_nat_rel m n qr" . | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 569 | then show ?thesis | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 570 | by (auto simp add: divmod_nat_def intro: theI elim: divmod_nat_rel_unique) | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 571 | qed | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 572 | |
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 573 | lemma divmod_nat_unique: | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 574 | assumes "divmod_nat_rel m n qr" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 575 | shows "divmod_nat m n = qr" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 576 | using assms by (auto intro: divmod_nat_rel_unique divmod_nat_rel_divmod_nat) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 577 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 578 | instantiation nat :: semiring_div | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 579 | begin | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 580 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 581 | definition div_nat where | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 582 | "m div n = fst (divmod_nat m n)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 583 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 584 | lemma fst_divmod_nat [simp]: | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 585 | "fst (divmod_nat m n) = m div n" | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 586 | by (simp add: div_nat_def) | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 587 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 588 | definition mod_nat where | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 589 | "m mod n = snd (divmod_nat m n)" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25162diff
changeset | 590 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 591 | lemma snd_divmod_nat [simp]: | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 592 | "snd (divmod_nat m n) = m mod n" | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 593 | by (simp add: mod_nat_def) | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 594 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 595 | lemma divmod_nat_div_mod: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 596 | "divmod_nat m n = (m div n, m mod n)" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 597 | by (simp add: prod_eq_iff) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 598 | |
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 599 | lemma div_nat_unique: | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 600 | assumes "divmod_nat_rel m n (q, r)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 601 | shows "m div n = q" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 602 | using assms by (auto dest!: divmod_nat_unique simp add: prod_eq_iff) | 
| 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 603 | |
| 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 604 | lemma mod_nat_unique: | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 605 | assumes "divmod_nat_rel m n (q, r)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 606 | shows "m mod n = r" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 607 | using assms by (auto dest!: divmod_nat_unique simp add: prod_eq_iff) | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25162diff
changeset | 608 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 609 | lemma divmod_nat_rel: "divmod_nat_rel m n (m div n, m mod n)" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 610 | using divmod_nat_rel_divmod_nat by (simp add: divmod_nat_div_mod) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 611 | |
| 47136 | 612 | lemma divmod_nat_zero: "divmod_nat m 0 = (0, m)" | 
| 613 | by (simp add: divmod_nat_unique divmod_nat_rel_def) | |
| 614 | ||
| 615 | lemma divmod_nat_zero_left: "divmod_nat 0 n = (0, 0)" | |
| 616 | by (simp add: divmod_nat_unique divmod_nat_rel_def) | |
| 25942 | 617 | |
| 47137 | 618 | lemma divmod_nat_base: "m < n \<Longrightarrow> divmod_nat m n = (0, m)" | 
| 619 | by (simp add: divmod_nat_unique divmod_nat_rel_def) | |
| 25942 | 620 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 621 | lemma divmod_nat_step: | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 622 | assumes "0 < n" and "n \<le> m" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 623 | shows "divmod_nat m n = (Suc ((m - n) div n), (m - n) mod n)" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 624 | proof (rule divmod_nat_unique) | 
| 47134 | 625 | have "divmod_nat_rel (m - n) n ((m - n) div n, (m - n) mod n)" | 
| 626 | by (rule divmod_nat_rel) | |
| 627 | thus "divmod_nat_rel m n (Suc ((m - n) div n), (m - n) mod n)" | |
| 628 | unfolding divmod_nat_rel_def using assms by auto | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 629 | qed | 
| 25942 | 630 | |
| 26300 | 631 | text {* The ''recursion'' equations for @{const div} and @{const mod} *}
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 632 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 633 | lemma div_less [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 634 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 635 | assumes "m < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 636 | shows "m div n = 0" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 637 | using assms divmod_nat_base by (simp add: prod_eq_iff) | 
| 25942 | 638 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 639 | lemma le_div_geq: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 640 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 641 | assumes "0 < n" and "n \<le> m" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 642 | shows "m div n = Suc ((m - n) div n)" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 643 | using assms divmod_nat_step by (simp add: prod_eq_iff) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 644 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 645 | lemma mod_less [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 646 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 647 | assumes "m < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 648 | shows "m mod n = m" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 649 | using assms divmod_nat_base by (simp add: prod_eq_iff) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 650 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 651 | lemma le_mod_geq: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 652 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 653 | assumes "n \<le> m" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 654 | shows "m mod n = (m - n) mod n" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 655 | using assms divmod_nat_step by (cases "n = 0") (simp_all add: prod_eq_iff) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 656 | |
| 47136 | 657 | instance proof | 
| 658 | fix m n :: nat | |
| 659 | show "m div n * n + m mod n = m" | |
| 660 | using divmod_nat_rel [of m n] by (simp add: divmod_nat_rel_def) | |
| 661 | next | |
| 662 | fix m n q :: nat | |
| 663 | assume "n \<noteq> 0" | |
| 664 | then show "(q + m * n) div n = m + q div n" | |
| 665 | by (induct m) (simp_all add: le_div_geq) | |
| 666 | next | |
| 667 | fix m n q :: nat | |
| 668 | assume "m \<noteq> 0" | |
| 669 | hence "\<And>a b. divmod_nat_rel n q (a, b) \<Longrightarrow> divmod_nat_rel (m * n) (m * q) (a, m * b)" | |
| 670 | unfolding divmod_nat_rel_def | |
| 671 | by (auto split: split_if_asm, simp_all add: algebra_simps) | |
| 672 | moreover from divmod_nat_rel have "divmod_nat_rel n q (n div q, n mod q)" . | |
| 673 | ultimately have "divmod_nat_rel (m * n) (m * q) (n div q, m * (n mod q))" . | |
| 674 | thus "(m * n) div (m * q) = n div q" by (rule div_nat_unique) | |
| 675 | next | |
| 676 | fix n :: nat show "n div 0 = 0" | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 677 | by (simp add: div_nat_def divmod_nat_zero) | 
| 47136 | 678 | next | 
| 679 | fix n :: nat show "0 div n = 0" | |
| 680 | by (simp add: div_nat_def divmod_nat_zero_left) | |
| 25942 | 681 | qed | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 682 | |
| 25942 | 683 | end | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 684 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 685 | lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 686 | let (q, r) = divmod_nat (m - n) n in (Suc q, r))" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 687 | by (simp add: prod_eq_iff prod_case_beta not_less le_div_geq le_mod_geq) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 688 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 689 | text {* Simproc for cancelling @{const div} and @{const mod} *}
 | 
| 25942 | 690 | |
| 30934 | 691 | ML {*
 | 
| 43594 | 692 | structure Cancel_Div_Mod_Nat = Cancel_Div_Mod | 
| 41550 | 693 | ( | 
| 30934 | 694 |   val div_name = @{const_name div};
 | 
| 695 |   val mod_name = @{const_name mod};
 | |
| 696 | val mk_binop = HOLogic.mk_binop; | |
| 48561 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 697 |   val mk_plus = HOLogic.mk_binop @{const_name Groups.plus};
 | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 698 |   val dest_plus = HOLogic.dest_bin @{const_name Groups.plus} HOLogic.natT;
 | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 699 | fun mk_sum [] = HOLogic.zero | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 700 | | mk_sum [t] = t | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 701 | | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 702 | fun dest_sum tm = | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 703 | if HOLogic.is_zero tm then [] | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 704 | else | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 705 | (case try HOLogic.dest_Suc tm of | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 706 | SOME t => HOLogic.Suc_zero :: dest_sum t | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 707 | | NONE => | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 708 | (case try dest_plus tm of | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 709 | SOME (t, u) => dest_sum t @ dest_sum u | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 710 | | NONE => [tm])); | 
| 25942 | 711 | |
| 30934 | 712 |   val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}];
 | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 713 | |
| 30934 | 714 | val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac | 
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34982diff
changeset | 715 |     (@{thm add_0_left} :: @{thm add_0_right} :: @{thms add_ac}))
 | 
| 41550 | 716 | ) | 
| 25942 | 717 | *} | 
| 718 | ||
| 43594 | 719 | simproc_setup cancel_div_mod_nat ("(m::nat) + n") = {* K Cancel_Div_Mod_Nat.proc *}
 | 
| 720 | ||
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 721 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 722 | subsubsection {* Quotient *}
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 723 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 724 | lemma div_geq: "0 < n \<Longrightarrow> \<not> m < n \<Longrightarrow> m div n = Suc ((m - n) div n)" | 
| 29667 | 725 | by (simp add: le_div_geq linorder_not_less) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 726 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 727 | lemma div_if: "0 < n \<Longrightarrow> m div n = (if m < n then 0 else Suc ((m - n) div n))" | 
| 29667 | 728 | by (simp add: div_geq) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 729 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 730 | lemma div_mult_self_is_m [simp]: "0<n ==> (m*n) div n = (m::nat)" | 
| 29667 | 731 | by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 732 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 733 | lemma div_mult_self1_is_m [simp]: "0<n ==> (n*m) div n = (m::nat)" | 
| 29667 | 734 | by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 735 | |
| 25942 | 736 | |
| 737 | subsubsection {* Remainder *}
 | |
| 738 | ||
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 739 | lemma mod_less_divisor [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 740 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 741 | assumes "n > 0" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 742 | shows "m mod n < (n::nat)" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 743 | using assms divmod_nat_rel [of m n] unfolding divmod_nat_rel_def by auto | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 744 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 745 | lemma mod_less_eq_dividend [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 746 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 747 | shows "m mod n \<le> m" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 748 | proof (rule add_leD2) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 749 | from mod_div_equality have "m div n * n + m mod n = m" . | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 750 | then show "m div n * n + m mod n \<le> m" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 751 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 752 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 753 | lemma mod_geq: "\<not> m < (n\<Colon>nat) \<Longrightarrow> m mod n = (m - n) mod n" | 
| 29667 | 754 | by (simp add: le_mod_geq linorder_not_less) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 755 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 756 | lemma mod_if: "m mod (n\<Colon>nat) = (if m < n then m else (m - n) mod n)" | 
| 29667 | 757 | by (simp add: le_mod_geq) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 758 | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 759 | lemma mod_1 [simp]: "m mod Suc 0 = 0" | 
| 29667 | 760 | by (induct m) (simp_all add: mod_geq) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 761 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 762 | (* a simple rearrangement of mod_div_equality: *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 763 | lemma mult_div_cancel: "(n::nat) * (m div n) = m - (m mod n)" | 
| 47138 | 764 | using mod_div_equality2 [of n m] by arith | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 765 | |
| 15439 | 766 | lemma mod_le_divisor[simp]: "0 < n \<Longrightarrow> m mod n \<le> (n::nat)" | 
| 22718 | 767 | apply (drule mod_less_divisor [where m = m]) | 
| 768 | apply simp | |
| 769 | done | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 770 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 771 | subsubsection {* Quotient and Remainder *}
 | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 772 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 773 | lemma divmod_nat_rel_mult1_eq: | 
| 46552 | 774 | "divmod_nat_rel b c (q, r) | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 775 | \<Longrightarrow> divmod_nat_rel (a * b) c (a * q + a * r div c, a * r mod c)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 776 | by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 777 | |
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 778 | lemma div_mult1_eq: | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 779 | "(a * b) div c = a * (b div c) + a * (b mod c) div (c::nat)" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 780 | by (blast intro: divmod_nat_rel_mult1_eq [THEN div_nat_unique] divmod_nat_rel) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 781 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 782 | lemma divmod_nat_rel_add1_eq: | 
| 46552 | 783 | "divmod_nat_rel a c (aq, ar) \<Longrightarrow> divmod_nat_rel b c (bq, br) | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 784 | \<Longrightarrow> divmod_nat_rel (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 785 | by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 786 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 787 | (*NOT suitable for rewriting: the RHS has an instance of the LHS*) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 788 | lemma div_add1_eq: | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 789 | "(a+b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 790 | by (blast intro: divmod_nat_rel_add1_eq [THEN div_nat_unique] divmod_nat_rel) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 791 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 792 | lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c" | 
| 22718 | 793 | apply (cut_tac m = q and n = c in mod_less_divisor) | 
| 794 | apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto) | |
| 795 | apply (erule_tac P = "%x. ?lhs < ?rhs x" in ssubst) | |
| 796 | apply (simp add: add_mult_distrib2) | |
| 797 | done | |
| 10559 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 798 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 799 | lemma divmod_nat_rel_mult2_eq: | 
| 46552 | 800 | "divmod_nat_rel a b (q, r) | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 801 | \<Longrightarrow> divmod_nat_rel a (b * c) (q div c, b *(q mod c) + r)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 802 | by (auto simp add: mult_ac divmod_nat_rel_def add_mult_distrib2 [symmetric] mod_lemma) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 803 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 804 | lemma div_mult2_eq: "a div (b*c) = (a div b) div (c::nat)" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 805 | by (force simp add: divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN div_nat_unique]) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 806 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 807 | lemma mod_mult2_eq: "a mod (b*c) = b*(a div b mod c) + a mod (b::nat)" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 808 | by (auto simp add: mult_commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_nat_unique]) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 809 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 810 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 811 | subsubsection {* Further Facts about Quotient and Remainder *}
 | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 812 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 813 | lemma div_1 [simp]: "m div Suc 0 = m" | 
| 29667 | 814 | by (induct m) (simp_all add: div_geq) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 815 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 816 | (* Monotonicity of div in first argument *) | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 817 | lemma div_le_mono [rule_format (no_asm)]: | 
| 22718 | 818 | "\<forall>m::nat. m \<le> n --> (m div k) \<le> (n div k)" | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 819 | apply (case_tac "k=0", simp) | 
| 15251 | 820 | apply (induct "n" rule: nat_less_induct, clarify) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 821 | apply (case_tac "n<k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 822 | (* 1 case n<k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 823 | apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 824 | (* 2 case n >= k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 825 | apply (case_tac "m<k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 826 | (* 2.1 case m<k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 827 | apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 828 | (* 2.2 case m>=k *) | 
| 15439 | 829 | apply (simp add: div_geq diff_le_mono) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 830 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 831 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 832 | (* Antimonotonicity of div in second argument *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 833 | lemma div_le_mono2: "!!m::nat. [| 0<m; m\<le>n |] ==> (k div n) \<le> (k div m)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 834 | apply (subgoal_tac "0<n") | 
| 22718 | 835 | prefer 2 apply simp | 
| 15251 | 836 | apply (induct_tac k rule: nat_less_induct) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 837 | apply (rename_tac "k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 838 | apply (case_tac "k<n", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 839 | apply (subgoal_tac "~ (k<m) ") | 
| 22718 | 840 | prefer 2 apply simp | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 841 | apply (simp add: div_geq) | 
| 15251 | 842 | apply (subgoal_tac "(k-n) div n \<le> (k-m) div n") | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 843 | prefer 2 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 844 | apply (blast intro: div_le_mono diff_le_mono2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 845 | apply (rule le_trans, simp) | 
| 15439 | 846 | apply (simp) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 847 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 848 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 849 | lemma div_le_dividend [simp]: "m div n \<le> (m::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 850 | apply (case_tac "n=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 851 | apply (subgoal_tac "m div n \<le> m div 1", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 852 | apply (rule div_le_mono2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 853 | apply (simp_all (no_asm_simp)) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 854 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 855 | |
| 22718 | 856 | (* Similar for "less than" *) | 
| 47138 | 857 | lemma div_less_dividend [simp]: | 
| 858 | "\<lbrakk>(1::nat) < n; 0 < m\<rbrakk> \<Longrightarrow> m div n < m" | |
| 859 | apply (induct m rule: nat_less_induct) | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 860 | apply (rename_tac "m") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 861 | apply (case_tac "m<n", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 862 | apply (subgoal_tac "0<n") | 
| 22718 | 863 | prefer 2 apply simp | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 864 | apply (simp add: div_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 865 | apply (case_tac "n<m") | 
| 15251 | 866 | apply (subgoal_tac "(m-n) div n < (m-n) ") | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 867 | apply (rule impI less_trans_Suc)+ | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 868 | apply assumption | 
| 15439 | 869 | apply (simp_all) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 870 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 871 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 872 | text{*A fact for the mutilated chess board*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 873 | lemma mod_Suc: "Suc(m) mod n = (if Suc(m mod n) = n then 0 else Suc(m mod n))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 874 | apply (case_tac "n=0", simp) | 
| 15251 | 875 | apply (induct "m" rule: nat_less_induct) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 876 | apply (case_tac "Suc (na) <n") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 877 | (* case Suc(na) < n *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 878 | apply (frule lessI [THEN less_trans], simp add: less_not_refl3) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 879 | (* case n \<le> Suc(na) *) | 
| 16796 | 880 | apply (simp add: linorder_not_less le_Suc_eq mod_geq) | 
| 15439 | 881 | apply (auto simp add: Suc_diff_le le_mod_geq) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 882 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 883 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 884 | lemma mod_eq_0_iff: "(m mod d = 0) = (\<exists>q::nat. m = d*q)" | 
| 29667 | 885 | by (auto simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) | 
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16796diff
changeset | 886 | |
| 22718 | 887 | lemmas mod_eq_0D [dest!] = mod_eq_0_iff [THEN iffD1] | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 888 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 889 | (*Loses information, namely we also have r<d provided d is nonzero*) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 890 | lemma mod_eqD: "(m mod d = r) ==> \<exists>q::nat. m = r + q*d" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 891 | apply (cut_tac a = m in mod_div_equality) | 
| 22718 | 892 | apply (simp only: add_ac) | 
| 893 | apply (blast intro: sym) | |
| 894 | done | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 895 | |
| 13152 | 896 | lemma split_div: | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 897 | "P(n div k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 898 | ((k = 0 \<longrightarrow> P 0) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P i)))" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 899 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 900 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 901 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 902 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 903 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 904 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 905 | with P show ?Q by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 906 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 907 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 908 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 909 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 910 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 911 | assume n: "n = k*i + j" and j: "j < k" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 912 | show "P i" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 913 | proof (cases) | 
| 22718 | 914 | assume "i = 0" | 
| 915 | with n j P show "P i" by simp | |
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 916 | next | 
| 22718 | 917 | assume "i \<noteq> 0" | 
| 918 | with not0 n j P show "P i" by(simp add:add_ac) | |
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 919 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 920 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 921 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 922 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 923 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 924 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 925 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 926 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 927 | with Q show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 928 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 929 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 930 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 931 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 932 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 933 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 934 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 935 | |
| 13882 | 936 | lemma split_div_lemma: | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 937 | assumes "0 < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 938 | shows "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> q = ((m\<Colon>nat) div n)" (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 939 | proof | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 940 | assume ?rhs | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 941 | with mult_div_cancel have nq: "n * q = m - (m mod n)" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 942 | then have A: "n * q \<le> m" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 943 | have "n - (m mod n) > 0" using mod_less_divisor assms by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 944 | then have "m < m + (n - (m mod n))" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 945 | then have "m < n + (m - (m mod n))" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 946 | with nq have "m < n + n * q" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 947 | then have B: "m < n * Suc q" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 948 | from A B show ?lhs .. | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 949 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 950 | assume P: ?lhs | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 951 | then have "divmod_nat_rel m n (q, m - n * q)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 952 | unfolding divmod_nat_rel_def by (auto simp add: mult_ac) | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 953 | with divmod_nat_rel_unique divmod_nat_rel [of m n] | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 954 | have "(q, m - n * q) = (m div n, m mod n)" by auto | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 955 | then show ?rhs by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 956 | qed | 
| 13882 | 957 | |
| 958 | theorem split_div': | |
| 959 | "P ((m::nat) div n) = ((n = 0 \<and> P 0) \<or> | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 960 | (\<exists>q. (n * q \<le> m \<and> m < n * (Suc q)) \<and> P q))" | 
| 13882 | 961 | apply (case_tac "0 < n") | 
| 962 | apply (simp only: add: split_div_lemma) | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 963 | apply simp_all | 
| 13882 | 964 | done | 
| 965 | ||
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 966 | lemma split_mod: | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 967 | "P(n mod k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 968 | ((k = 0 \<longrightarrow> P n) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P j)))" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 969 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 970 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 971 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 972 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 973 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 974 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 975 | with P show ?Q by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 976 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 977 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 978 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 979 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 980 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 981 | assume "n = k*i + j" "j < k" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 982 | thus "P j" using not0 P by(simp add:add_ac mult_ac) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 983 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 984 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 985 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 986 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 987 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 988 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 989 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 990 | with Q show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 991 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 992 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 993 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 994 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 995 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 996 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 997 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 998 | |
| 13882 | 999 | theorem mod_div_equality': "(m::nat) mod n = m - (m div n) * n" | 
| 47138 | 1000 | using mod_div_equality [of m n] by arith | 
| 1001 | ||
| 1002 | lemma div_mod_equality': "(m::nat) div n * n = m - m mod n" | |
| 1003 | using mod_div_equality [of m n] by arith | |
| 1004 | (* FIXME: very similar to mult_div_cancel *) | |
| 22800 | 1005 | |
| 1006 | ||
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1007 | subsubsection {* An ``induction'' law for modulus arithmetic. *}
 | 
| 14640 | 1008 | |
| 1009 | lemma mod_induct_0: | |
| 1010 | assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)" | |
| 1011 | and base: "P i" and i: "i<p" | |
| 1012 | shows "P 0" | |
| 1013 | proof (rule ccontr) | |
| 1014 | assume contra: "\<not>(P 0)" | |
| 1015 | from i have p: "0<p" by simp | |
| 1016 | have "\<forall>k. 0<k \<longrightarrow> \<not> P (p-k)" (is "\<forall>k. ?A k") | |
| 1017 | proof | |
| 1018 | fix k | |
| 1019 | show "?A k" | |
| 1020 | proof (induct k) | |
| 1021 | show "?A 0" by simp -- "by contradiction" | |
| 1022 | next | |
| 1023 | fix n | |
| 1024 | assume ih: "?A n" | |
| 1025 | show "?A (Suc n)" | |
| 1026 | proof (clarsimp) | |
| 22718 | 1027 | assume y: "P (p - Suc n)" | 
| 1028 | have n: "Suc n < p" | |
| 1029 | proof (rule ccontr) | |
| 1030 | assume "\<not>(Suc n < p)" | |
| 1031 | hence "p - Suc n = 0" | |
| 1032 | by simp | |
| 1033 | with y contra show "False" | |
| 1034 | by simp | |
| 1035 | qed | |
| 1036 | hence n2: "Suc (p - Suc n) = p-n" by arith | |
| 1037 | from p have "p - Suc n < p" by arith | |
| 1038 | with y step have z: "P ((Suc (p - Suc n)) mod p)" | |
| 1039 | by blast | |
| 1040 | show "False" | |
| 1041 | proof (cases "n=0") | |
| 1042 | case True | |
| 1043 | with z n2 contra show ?thesis by simp | |
| 1044 | next | |
| 1045 | case False | |
| 1046 | with p have "p-n < p" by arith | |
| 1047 | with z n2 False ih show ?thesis by simp | |
| 1048 | qed | |
| 14640 | 1049 | qed | 
| 1050 | qed | |
| 1051 | qed | |
| 1052 | moreover | |
| 1053 | from i obtain k where "0<k \<and> i+k=p" | |
| 1054 | by (blast dest: less_imp_add_positive) | |
| 1055 | hence "0<k \<and> i=p-k" by auto | |
| 1056 | moreover | |
| 1057 | note base | |
| 1058 | ultimately | |
| 1059 | show "False" by blast | |
| 1060 | qed | |
| 1061 | ||
| 1062 | lemma mod_induct: | |
| 1063 | assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)" | |
| 1064 | and base: "P i" and i: "i<p" and j: "j<p" | |
| 1065 | shows "P j" | |
| 1066 | proof - | |
| 1067 | have "\<forall>j<p. P j" | |
| 1068 | proof | |
| 1069 | fix j | |
| 1070 | show "j<p \<longrightarrow> P j" (is "?A j") | |
| 1071 | proof (induct j) | |
| 1072 | from step base i show "?A 0" | |
| 22718 | 1073 | by (auto elim: mod_induct_0) | 
| 14640 | 1074 | next | 
| 1075 | fix k | |
| 1076 | assume ih: "?A k" | |
| 1077 | show "?A (Suc k)" | |
| 1078 | proof | |
| 22718 | 1079 | assume suc: "Suc k < p" | 
| 1080 | hence k: "k<p" by simp | |
| 1081 | with ih have "P k" .. | |
| 1082 | with step k have "P (Suc k mod p)" | |
| 1083 | by blast | |
| 1084 | moreover | |
| 1085 | from suc have "Suc k mod p = Suc k" | |
| 1086 | by simp | |
| 1087 | ultimately | |
| 1088 | show "P (Suc k)" by simp | |
| 14640 | 1089 | qed | 
| 1090 | qed | |
| 1091 | qed | |
| 1092 | with j show ?thesis by blast | |
| 1093 | qed | |
| 1094 | ||
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1095 | lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)" | 
| 47138 | 1096 | by (simp add: numeral_2_eq_2 le_div_geq) | 
| 1097 | ||
| 1098 | lemma mod2_Suc_Suc [simp]: "Suc (Suc m) mod 2 = m mod 2" | |
| 1099 | by (simp add: numeral_2_eq_2 le_mod_geq) | |
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1100 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1101 | lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)" | 
| 47217 
501b9bbd0d6e
removed redundant nat-specific copies of theorems
 huffman parents: 
47167diff
changeset | 1102 | by (simp add: mult_2 [symmetric]) | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1103 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1104 | lemma mod2_gr_0 [simp]: "0 < (m\<Colon>nat) mod 2 \<longleftrightarrow> m mod 2 = 1" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1105 | proof - | 
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1106 |   { fix n :: nat have  "(n::nat) < 2 \<Longrightarrow> n = 0 \<or> n = 1" by (cases n) simp_all }
 | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1107 | moreover have "m mod 2 < 2" by simp | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1108 | ultimately have "m mod 2 = 0 \<or> m mod 2 = 1" . | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1109 | then show ?thesis by auto | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1110 | qed | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1111 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1112 | text{*These lemmas collapse some needless occurrences of Suc:
 | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1113 | at least three Sucs, since two and fewer are rewritten back to Suc again! | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1114 | We already have some rules to simplify operands smaller than 3.*} | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1115 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1116 | lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1117 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1118 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1119 | lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1120 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1121 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1122 | lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1123 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1124 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1125 | lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1126 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1127 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1128 | lemmas Suc_div_eq_add3_div_numeral [simp] = Suc_div_eq_add3_div [of _ "numeral v"] for v | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1129 | lemmas Suc_mod_eq_add3_mod_numeral [simp] = Suc_mod_eq_add3_mod [of _ "numeral v"] for v | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1130 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1131 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1132 | lemma Suc_times_mod_eq: "1<k ==> Suc (k * m) mod k = 1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1133 | apply (induct "m") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1134 | apply (simp_all add: mod_Suc) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1135 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1136 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1137 | declare Suc_times_mod_eq [of "numeral w", simp] for w | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1138 | |
| 47138 | 1139 | lemma Suc_div_le_mono [simp]: "n div k \<le> (Suc n) div k" | 
| 1140 | by (simp add: div_le_mono) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1141 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1142 | lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1143 | by (cases n) simp_all | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1144 | |
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1145 | lemma div_2_gt_zero [simp]: assumes A: "(1::nat) < n" shows "0 < n div 2" | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1146 | proof - | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1147 | from A have B: "0 < n - 1" and C: "n - 1 + 1 = n" by simp_all | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1148 | from Suc_n_div_2_gt_zero [OF B] C show ?thesis by simp | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1149 | qed | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1150 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1151 | (* Potential use of algebra : Equality modulo n*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1152 | lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1153 | by (simp add: mult_ac add_ac) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1154 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1155 | lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1156 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1157 | have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1158 | also have "... = Suc m mod n" by (rule mod_mult_self3) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1159 | finally show ?thesis . | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1160 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1161 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1162 | lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1163 | apply (subst mod_Suc [of m]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1164 | apply (subst mod_Suc [of "m mod n"], simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1165 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1166 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1167 | lemma mod_2_not_eq_zero_eq_one_nat: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1168 | fixes n :: nat | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1169 | shows "n mod 2 \<noteq> 0 \<longleftrightarrow> n mod 2 = 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1170 | by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1171 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1172 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1173 | subsection {* Division on @{typ int} *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1174 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1175 | definition divmod_int_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1176 |     --{*definition of quotient and remainder*}
 | 
| 47139 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1177 | "divmod_int_rel a b = (\<lambda>(q, r). a = b * q + r \<and> | 
| 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1178 | (if 0 < b then 0 \<le> r \<and> r < b else if b < 0 then b < r \<and> r \<le> 0 else q = 0))" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1179 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1180 | definition adjust :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1181 |     --{*for the division algorithm*}
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1182 | "adjust b = (\<lambda>(q, r). if 0 \<le> r - b then (2 * q + 1, r - b) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1183 | else (2 * q, r))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1184 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1185 | text{*algorithm for the case @{text "a\<ge>0, b>0"}*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1186 | function posDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1187 | "posDivAlg a b = (if a < b \<or> b \<le> 0 then (0, a) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1188 | else adjust b (posDivAlg a (2 * b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1189 | by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1190 | termination by (relation "measure (\<lambda>(a, b). nat (a - b + 1))") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1191 | (auto simp add: mult_2) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1192 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1193 | text{*algorithm for the case @{text "a<0, b>0"}*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1194 | function negDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1195 | "negDivAlg a b = (if 0 \<le>a + b \<or> b \<le> 0 then (-1, a + b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1196 | else adjust b (negDivAlg a (2 * b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1197 | by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1198 | termination by (relation "measure (\<lambda>(a, b). nat (- a - b))") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1199 | (auto simp add: mult_2) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1200 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1201 | text{*algorithm for the general case @{term "b\<noteq>0"}*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1202 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1203 | definition divmod_int :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1204 |     --{*The full division algorithm considers all possible signs for a, b
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1205 |        including the special case @{text "a=0, b<0"} because 
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1206 |        @{term negDivAlg} requires @{term "a<0"}.*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1207 | "divmod_int a b = (if 0 \<le> a then if 0 \<le> b then posDivAlg a b | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1208 | else if a = 0 then (0, 0) | 
| 46560 
8e252a608765
remove constant negateSnd in favor of 'apsnd uminus' (from Florian Haftmann)
 huffman parents: 
46552diff
changeset | 1209 | else apsnd uminus (negDivAlg (-a) (-b)) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1210 | else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1211 | if 0 < b then negDivAlg a b | 
| 46560 
8e252a608765
remove constant negateSnd in favor of 'apsnd uminus' (from Florian Haftmann)
 huffman parents: 
46552diff
changeset | 1212 | else apsnd uminus (posDivAlg (-a) (-b)))" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1213 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1214 | instantiation int :: Divides.div | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1215 | begin | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1216 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1217 | definition div_int where | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1218 | "a div b = fst (divmod_int a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1219 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1220 | lemma fst_divmod_int [simp]: | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1221 | "fst (divmod_int a b) = a div b" | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1222 | by (simp add: div_int_def) | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1223 | |
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1224 | definition mod_int where | 
| 46560 
8e252a608765
remove constant negateSnd in favor of 'apsnd uminus' (from Florian Haftmann)
 huffman parents: 
46552diff
changeset | 1225 | "a mod b = snd (divmod_int a b)" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1226 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1227 | lemma snd_divmod_int [simp]: | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1228 | "snd (divmod_int a b) = a mod b" | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1229 | by (simp add: mod_int_def) | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1230 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1231 | instance .. | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1232 | |
| 3366 | 1233 | end | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1234 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1235 | lemma divmod_int_mod_div: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1236 | "divmod_int p q = (p div q, p mod q)" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1237 | by (simp add: prod_eq_iff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1238 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1239 | text{*
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1240 | Here is the division algorithm in ML: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1241 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1242 | \begin{verbatim}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1243 | fun posDivAlg (a,b) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1244 | if a<b then (0,a) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1245 | else let val (q,r) = posDivAlg(a, 2*b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1246 | in if 0\<le>r-b then (2*q+1, r-b) else (2*q, r) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1247 | end | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1248 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1249 | fun negDivAlg (a,b) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1250 | if 0\<le>a+b then (~1,a+b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1251 | else let val (q,r) = negDivAlg(a, 2*b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1252 | in if 0\<le>r-b then (2*q+1, r-b) else (2*q, r) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1253 | end; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1254 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1255 | fun negateSnd (q,r:int) = (q,~r); | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1256 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1257 | fun divmod (a,b) = if 0\<le>a then | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1258 | if b>0 then posDivAlg (a,b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1259 | else if a=0 then (0,0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1260 | else negateSnd (negDivAlg (~a,~b)) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1261 | else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1262 | if 0<b then negDivAlg (a,b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1263 | else negateSnd (posDivAlg (~a,~b)); | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1264 | \end{verbatim}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1265 | *} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1266 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1267 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1268 | subsubsection {* Uniqueness and Monotonicity of Quotients and Remainders *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1269 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1270 | lemma unique_quotient_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1271 | "[| b*q' + r' \<le> b*q + r; 0 \<le> r'; r' < b; r < b |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1272 | ==> q' \<le> (q::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1273 | apply (subgoal_tac "r' + b * (q'-q) \<le> r") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1274 | prefer 2 apply (simp add: right_diff_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1275 | apply (subgoal_tac "0 < b * (1 + q - q') ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1276 | apply (erule_tac [2] order_le_less_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1277 | prefer 2 apply (simp add: right_diff_distrib right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1278 | apply (subgoal_tac "b * q' < b * (1 + q) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1279 | prefer 2 apply (simp add: right_diff_distrib right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1280 | apply (simp add: mult_less_cancel_left) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1281 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1282 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1283 | lemma unique_quotient_lemma_neg: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1284 | "[| b*q' + r' \<le> b*q + r; r \<le> 0; b < r; b < r' |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1285 | ==> q \<le> (q'::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1286 | by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1287 | auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1288 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1289 | lemma unique_quotient: | 
| 46552 | 1290 | "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r') |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1291 | ==> q = q'" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1292 | apply (simp add: divmod_int_rel_def linorder_neq_iff split: split_if_asm) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1293 | apply (blast intro: order_antisym | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1294 | dest: order_eq_refl [THEN unique_quotient_lemma] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1295 | order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1296 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1297 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1298 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1299 | lemma unique_remainder: | 
| 46552 | 1300 | "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r') |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1301 | ==> r = r'" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1302 | apply (subgoal_tac "q = q'") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1303 | apply (simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1304 | apply (blast intro: unique_quotient) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1305 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1306 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1307 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1308 | subsubsection {* Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1309 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1310 | text{*And positive divisors*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1311 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1312 | lemma adjust_eq [simp]: | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1313 | "adjust b (q, r) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1314 | (let diff = r - b in | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1315 | if 0 \<le> diff then (2 * q + 1, diff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1316 | else (2*q, r))" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1317 | by (simp add: Let_def adjust_def) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1318 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1319 | declare posDivAlg.simps [simp del] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1320 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1321 | text{*use with a simproc to avoid repeatedly proving the premise*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1322 | lemma posDivAlg_eqn: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1323 | "0 < b ==> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1324 | posDivAlg a b = (if a<b then (0,a) else adjust b (posDivAlg a (2*b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1325 | by (rule posDivAlg.simps [THEN trans], simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1326 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1327 | text{*Correctness of @{term posDivAlg}: it computes quotients correctly*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1328 | theorem posDivAlg_correct: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1329 | assumes "0 \<le> a" and "0 < b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1330 | shows "divmod_int_rel a b (posDivAlg a b)" | 
| 41550 | 1331 | using assms | 
| 1332 | apply (induct a b rule: posDivAlg.induct) | |
| 1333 | apply auto | |
| 1334 | apply (simp add: divmod_int_rel_def) | |
| 1335 | apply (subst posDivAlg_eqn, simp add: right_distrib) | |
| 1336 | apply (case_tac "a < b") | |
| 1337 | apply simp_all | |
| 1338 | apply (erule splitE) | |
| 1339 | apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) | |
| 1340 | done | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1341 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1342 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1343 | subsubsection {* Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1344 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1345 | text{*And positive divisors*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1346 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1347 | declare negDivAlg.simps [simp del] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1348 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1349 | text{*use with a simproc to avoid repeatedly proving the premise*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1350 | lemma negDivAlg_eqn: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1351 | "0 < b ==> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1352 | negDivAlg a b = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1353 | (if 0\<le>a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1354 | by (rule negDivAlg.simps [THEN trans], simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1355 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1356 | (*Correctness of negDivAlg: it computes quotients correctly | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1357 | It doesn't work if a=0 because the 0/b equals 0, not -1*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1358 | lemma negDivAlg_correct: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1359 | assumes "a < 0" and "b > 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1360 | shows "divmod_int_rel a b (negDivAlg a b)" | 
| 41550 | 1361 | using assms | 
| 1362 | apply (induct a b rule: negDivAlg.induct) | |
| 1363 | apply (auto simp add: linorder_not_le) | |
| 1364 | apply (simp add: divmod_int_rel_def) | |
| 1365 | apply (subst negDivAlg_eqn, assumption) | |
| 1366 | apply (case_tac "a + b < (0\<Colon>int)") | |
| 1367 | apply simp_all | |
| 1368 | apply (erule splitE) | |
| 1369 | apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) | |
| 1370 | done | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1371 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1372 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1373 | subsubsection {* Existence Shown by Proving the Division Algorithm to be Correct *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1374 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1375 | (*the case a=0*) | 
| 47139 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1376 | lemma divmod_int_rel_0: "divmod_int_rel 0 b (0, 0)" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1377 | by (auto simp add: divmod_int_rel_def linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1378 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1379 | lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1380 | by (subst posDivAlg.simps, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1381 | |
| 47139 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1382 | lemma posDivAlg_0_right [simp]: "posDivAlg a 0 = (0, a)" | 
| 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1383 | by (subst posDivAlg.simps, auto) | 
| 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1384 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1385 | lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1386 | by (subst negDivAlg.simps, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1387 | |
| 46560 
8e252a608765
remove constant negateSnd in favor of 'apsnd uminus' (from Florian Haftmann)
 huffman parents: 
46552diff
changeset | 1388 | lemma divmod_int_rel_neg: "divmod_int_rel (-a) (-b) qr ==> divmod_int_rel a b (apsnd uminus qr)" | 
| 47139 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1389 | by (auto simp add: divmod_int_rel_def) | 
| 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1390 | |
| 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1391 | lemma divmod_int_correct: "divmod_int_rel a b (divmod_int a b)" | 
| 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1392 | apply (cases "b = 0", simp add: divmod_int_def divmod_int_rel_def) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1393 | by (force simp add: linorder_neq_iff divmod_int_rel_0 divmod_int_def divmod_int_rel_neg | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1394 | posDivAlg_correct negDivAlg_correct) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1395 | |
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1396 | lemma divmod_int_unique: | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1397 | assumes "divmod_int_rel a b qr" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1398 | shows "divmod_int a b = qr" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1399 | using assms divmod_int_correct [of a b] | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1400 | using unique_quotient [of a b] unique_remainder [of a b] | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1401 | by (metis pair_collapse) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1402 | |
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1403 | lemma divmod_int_rel_div_mod: "divmod_int_rel a b (a div b, a mod b)" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1404 | using divmod_int_correct by (simp add: divmod_int_mod_div) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1405 | |
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1406 | lemma div_int_unique: "divmod_int_rel a b (q, r) \<Longrightarrow> a div b = q" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1407 | by (simp add: divmod_int_rel_div_mod [THEN unique_quotient]) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1408 | |
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1409 | lemma mod_int_unique: "divmod_int_rel a b (q, r) \<Longrightarrow> a mod b = r" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1410 | by (simp add: divmod_int_rel_div_mod [THEN unique_remainder]) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1411 | |
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1412 | instance int :: ring_div | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1413 | proof | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1414 | fix a b :: int | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1415 | show "a div b * b + a mod b = a" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1416 | using divmod_int_rel_div_mod [of a b] | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1417 | unfolding divmod_int_rel_def by (simp add: mult_commute) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1418 | next | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1419 | fix a b c :: int | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1420 | assume "b \<noteq> 0" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1421 | hence "divmod_int_rel (a + c * b) b (c + a div b, a mod b)" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1422 | using divmod_int_rel_div_mod [of a b] | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1423 | unfolding divmod_int_rel_def by (auto simp: algebra_simps) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1424 | thus "(a + c * b) div b = c + a div b" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1425 | by (rule div_int_unique) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1426 | next | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1427 | fix a b c :: int | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1428 | assume "c \<noteq> 0" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1429 | hence "\<And>q r. divmod_int_rel a b (q, r) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1430 | \<Longrightarrow> divmod_int_rel (c * a) (c * b) (q, c * r)" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1431 | unfolding divmod_int_rel_def | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1432 | by - (rule linorder_cases [of 0 b], auto simp: algebra_simps | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1433 | mult_less_0_iff zero_less_mult_iff mult_strict_right_mono | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1434 | mult_strict_right_mono_neg zero_le_mult_iff mult_le_0_iff) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1435 | hence "divmod_int_rel (c * a) (c * b) (a div b, c * (a mod b))" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1436 | using divmod_int_rel_div_mod [of a b] . | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1437 | thus "(c * a) div (c * b) = a div b" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1438 | by (rule div_int_unique) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1439 | next | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1440 | fix a :: int show "a div 0 = 0" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1441 | by (rule div_int_unique, simp add: divmod_int_rel_def) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1442 | next | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1443 | fix a :: int show "0 div a = 0" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1444 | by (rule div_int_unique, auto simp add: divmod_int_rel_def) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1445 | qed | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1446 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1447 | text{*Basic laws about division and remainder*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1448 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1449 | lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)" | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1450 | by (fact mod_div_equality2 [symmetric]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1451 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1452 | text {* Tool setup *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1453 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1454 | (* FIXME: Theorem list add_0s doesn't exist, because Numeral0 has gone. *) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1455 | lemmas add_0s = add_0_left add_0_right | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1456 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1457 | ML {*
 | 
| 43594 | 1458 | structure Cancel_Div_Mod_Int = Cancel_Div_Mod | 
| 41550 | 1459 | ( | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1460 |   val div_name = @{const_name div};
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1461 |   val mod_name = @{const_name mod};
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1462 | val mk_binop = HOLogic.mk_binop; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1463 | val mk_sum = Arith_Data.mk_sum HOLogic.intT; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1464 | val dest_sum = Arith_Data.dest_sum; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1465 | |
| 47165 | 1466 |   val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}];
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1467 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1468 | val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1469 |     (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac}))
 | 
| 41550 | 1470 | ) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1471 | *} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1472 | |
| 43594 | 1473 | simproc_setup cancel_div_mod_int ("(k::int) + l") = {* K Cancel_Div_Mod_Int.proc *}
 | 
| 1474 | ||
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1475 | lemma pos_mod_conj: "(0::int) < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1476 | using divmod_int_correct [of a b] | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1477 | by (auto simp add: divmod_int_rel_def prod_eq_iff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1478 | |
| 45607 | 1479 | lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1] | 
| 1480 | and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2] | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1481 | |
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1482 | lemma neg_mod_conj: "b < (0::int) \<Longrightarrow> a mod b \<le> 0 \<and> b < a mod b" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1483 | using divmod_int_correct [of a b] | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1484 | by (auto simp add: divmod_int_rel_def prod_eq_iff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1485 | |
| 45607 | 1486 | lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1] | 
| 1487 | and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2] | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1488 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1489 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1490 | subsubsection {* General Properties of div and mod *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1491 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1492 | lemma div_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a div b = 0" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1493 | apply (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1494 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1495 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1496 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1497 | lemma div_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a div b = 0" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1498 | apply (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1499 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1500 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1501 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1502 | lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a div b = -1" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1503 | apply (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1504 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1505 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1506 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1507 | (*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1508 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1509 | lemma mod_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a mod b = a" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1510 | apply (rule_tac q = 0 in mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1511 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1512 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1513 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1514 | lemma mod_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a mod b = a" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1515 | apply (rule_tac q = 0 in mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1516 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1517 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1518 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1519 | lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a mod b = a+b" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1520 | apply (rule_tac q = "-1" in mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1521 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1522 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1523 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1524 | text{*There is no @{text mod_neg_pos_trivial}.*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1525 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1526 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1527 | subsubsection {* Laws for div and mod with Unary Minus *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1528 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1529 | lemma zminus1_lemma: | 
| 47139 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1530 | "divmod_int_rel a b (q, r) ==> b \<noteq> 0 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1531 | ==> divmod_int_rel (-a) b (if r=0 then -q else -q - 1, | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1532 | if r=0 then 0 else b-r)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1533 | by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_diff_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1534 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1535 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1536 | lemma zdiv_zminus1_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1537 | "b \<noteq> (0::int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1538 | ==> (-a) div b = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1539 | (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1540 | by (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN div_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1541 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1542 | lemma zmod_zminus1_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1543 | "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1544 | apply (case_tac "b = 0", simp) | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1545 | apply (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN mod_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1546 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1547 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1548 | lemma zmod_zminus1_not_zero: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1549 | fixes k l :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1550 | shows "- k mod l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1551 | unfolding zmod_zminus1_eq_if by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1552 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1553 | lemma zdiv_zminus2_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1554 | "b \<noteq> (0::int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1555 | ==> a div (-b) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1556 | (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | 
| 47159 | 1557 | by (simp add: zdiv_zminus1_eq_if div_minus_right) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1558 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1559 | lemma zmod_zminus2_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1560 | "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" | 
| 47159 | 1561 | by (simp add: zmod_zminus1_eq_if mod_minus_right) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1562 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1563 | lemma zmod_zminus2_not_zero: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1564 | fixes k l :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1565 | shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1566 | unfolding zmod_zminus2_eq_if by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1567 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1568 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1569 | subsubsection {* Computation of Division and Remainder *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1570 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1571 | lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1572 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1573 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1574 | lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1575 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1576 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1577 | text{*a positive, b positive *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1578 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1579 | lemma div_pos_pos: "[| 0 < a; 0 \<le> b |] ==> a div b = fst (posDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1580 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1581 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1582 | lemma mod_pos_pos: "[| 0 < a; 0 \<le> b |] ==> a mod b = snd (posDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1583 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1584 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1585 | text{*a negative, b positive *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1586 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1587 | lemma div_neg_pos: "[| a < 0; 0 < b |] ==> a div b = fst (negDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1588 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1589 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1590 | lemma mod_neg_pos: "[| a < 0; 0 < b |] ==> a mod b = snd (negDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1591 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1592 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1593 | text{*a positive, b negative *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1594 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1595 | lemma div_pos_neg: | 
| 46560 
8e252a608765
remove constant negateSnd in favor of 'apsnd uminus' (from Florian Haftmann)
 huffman parents: 
46552diff
changeset | 1596 | "[| 0 < a; b < 0 |] ==> a div b = fst (apsnd uminus (negDivAlg (-a) (-b)))" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1597 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1598 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1599 | lemma mod_pos_neg: | 
| 46560 
8e252a608765
remove constant negateSnd in favor of 'apsnd uminus' (from Florian Haftmann)
 huffman parents: 
46552diff
changeset | 1600 | "[| 0 < a; b < 0 |] ==> a mod b = snd (apsnd uminus (negDivAlg (-a) (-b)))" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1601 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1602 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1603 | text{*a negative, b negative *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1604 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1605 | lemma div_neg_neg: | 
| 46560 
8e252a608765
remove constant negateSnd in favor of 'apsnd uminus' (from Florian Haftmann)
 huffman parents: 
46552diff
changeset | 1606 | "[| a < 0; b \<le> 0 |] ==> a div b = fst (apsnd uminus (posDivAlg (-a) (-b)))" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1607 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1608 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1609 | lemma mod_neg_neg: | 
| 46560 
8e252a608765
remove constant negateSnd in favor of 'apsnd uminus' (from Florian Haftmann)
 huffman parents: 
46552diff
changeset | 1610 | "[| a < 0; b \<le> 0 |] ==> a mod b = snd (apsnd uminus (posDivAlg (-a) (-b)))" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1611 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1612 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1613 | text {*Simplify expresions in which div and mod combine numerical constants*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1614 | |
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1615 | lemma int_div_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a div b = q" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1616 | by (rule div_int_unique [of a b q r]) (simp add: divmod_int_rel_def) | 
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1617 | |
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1618 | lemma int_div_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a div b = q" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1619 | by (rule div_int_unique [of a b q r], | 
| 46552 | 1620 | simp add: divmod_int_rel_def) | 
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1621 | |
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1622 | lemma int_mod_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a mod b = r" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1623 | by (rule mod_int_unique [of a b q r], | 
| 46552 | 1624 | simp add: divmod_int_rel_def) | 
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1625 | |
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1626 | lemma int_mod_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a mod b = r" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1627 | by (rule mod_int_unique [of a b q r], | 
| 46552 | 1628 | simp add: divmod_int_rel_def) | 
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1629 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1630 | (* simprocs adapted from HOL/ex/Binary.thy *) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1631 | ML {*
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1632 | local | 
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1633 | val mk_number = HOLogic.mk_number HOLogic.intT | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1634 |   val plus = @{term "plus :: int \<Rightarrow> int \<Rightarrow> int"}
 | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1635 |   val times = @{term "times :: int \<Rightarrow> int \<Rightarrow> int"}
 | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1636 |   val zero = @{term "0 :: int"}
 | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1637 |   val less = @{term "op < :: int \<Rightarrow> int \<Rightarrow> bool"}
 | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1638 |   val le = @{term "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"}
 | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1639 |   val simps = @{thms arith_simps} @ @{thms rel_simps} @
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1640 |     map (fn th => th RS sym) [@{thm numeral_1_eq_1}]
 | 
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1641 | fun prove ctxt goal = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop goal) | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1642 | (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps simps)))); | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1643 | fun binary_proc proc ss ct = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1644 | (case Thm.term_of ct of | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1645 | _ $ t $ u => | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1646 | (case try (pairself (`(snd o HOLogic.dest_number))) (t, u) of | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1647 | SOME args => proc (Simplifier.the_context ss) args | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1648 | | NONE => NONE) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1649 | | _ => NONE); | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1650 | in | 
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1651 | fun divmod_proc posrule negrule = | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1652 | binary_proc (fn ctxt => fn ((a, t), (b, u)) => | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1653 | if b = 0 then NONE else let | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1654 | val (q, r) = pairself mk_number (Integer.div_mod a b) | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1655 | val goal1 = HOLogic.mk_eq (t, plus $ (times $ u $ q) $ r) | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1656 | val (goal2, goal3, rule) = if b > 0 | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1657 | then (le $ zero $ r, less $ r $ u, posrule RS eq_reflection) | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1658 | else (le $ r $ zero, less $ u $ r, negrule RS eq_reflection) | 
| 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1659 | in SOME (rule OF map (prove ctxt) [goal1, goal2, goal3]) end) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1660 | end | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1661 | *} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1662 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1663 | simproc_setup binary_int_div | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1664 |   ("numeral m div numeral n :: int" |
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1665 | "numeral m div neg_numeral n :: int" | | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1666 | "neg_numeral m div numeral n :: int" | | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1667 | "neg_numeral m div neg_numeral n :: int") = | 
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1668 |   {* K (divmod_proc @{thm int_div_pos_eq} @{thm int_div_neg_eq}) *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1669 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1670 | simproc_setup binary_int_mod | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1671 |   ("numeral m mod numeral n :: int" |
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1672 | "numeral m mod neg_numeral n :: int" | | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1673 | "neg_numeral m mod numeral n :: int" | | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1674 | "neg_numeral m mod neg_numeral n :: int") = | 
| 45530 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 huffman parents: 
45231diff
changeset | 1675 |   {* K (divmod_proc @{thm int_mod_pos_eq} @{thm int_mod_neg_eq}) *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1676 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1677 | lemmas posDivAlg_eqn_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1678 | posDivAlg_eqn [of "numeral v" "numeral w", OF zero_less_numeral] for v w | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1679 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1680 | lemmas negDivAlg_eqn_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1681 | negDivAlg_eqn [of "numeral v" "neg_numeral w", OF zero_less_numeral] for v w | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1682 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1683 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1684 | text{*Special-case simplification *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1685 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1686 | (** The last remaining special cases for constant arithmetic: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1687 | 1 div z and 1 mod z **) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1688 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1689 | lemmas div_pos_pos_1_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1690 | div_pos_pos [OF zero_less_one, of "numeral w", OF zero_le_numeral] for w | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1691 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1692 | lemmas div_pos_neg_1_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1693 | div_pos_neg [OF zero_less_one, of "neg_numeral w", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1694 | OF neg_numeral_less_zero] for w | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1695 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1696 | lemmas mod_pos_pos_1_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1697 | mod_pos_pos [OF zero_less_one, of "numeral w", OF zero_le_numeral] for w | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1698 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1699 | lemmas mod_pos_neg_1_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1700 | mod_pos_neg [OF zero_less_one, of "neg_numeral w", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1701 | OF neg_numeral_less_zero] for w | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1702 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1703 | lemmas posDivAlg_eqn_1_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1704 | posDivAlg_eqn [of concl: 1 "numeral w", OF zero_less_numeral] for w | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1705 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1706 | lemmas negDivAlg_eqn_1_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1707 | negDivAlg_eqn [of concl: 1 "numeral w", OF zero_less_numeral] for w | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1708 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1709 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1710 | subsubsection {* Monotonicity in the First Argument (Dividend) *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1711 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1712 | lemma zdiv_mono1: "[| a \<le> a'; 0 < (b::int) |] ==> a div b \<le> a' div b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1713 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1714 | apply (cut_tac a = a' and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1715 | apply (rule unique_quotient_lemma) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1716 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1717 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1718 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1719 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1720 | lemma zdiv_mono1_neg: "[| a \<le> a'; (b::int) < 0 |] ==> a' div b \<le> a div b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1721 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1722 | apply (cut_tac a = a' and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1723 | apply (rule unique_quotient_lemma_neg) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1724 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1725 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1726 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1727 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1728 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1729 | subsubsection {* Monotonicity in the Second Argument (Divisor) *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1730 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1731 | lemma q_pos_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1732 | "[| 0 \<le> b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \<le> (q'::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1733 | apply (subgoal_tac "0 < b'* (q' + 1) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1734 | apply (simp add: zero_less_mult_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1735 | apply (simp add: right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1736 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1737 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1738 | lemma zdiv_mono2_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1739 | "[| b*q + r = b'*q' + r'; 0 \<le> b'*q' + r'; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1740 | r' < b'; 0 \<le> r; 0 < b'; b' \<le> b |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1741 | ==> q \<le> (q'::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1742 | apply (frule q_pos_lemma, assumption+) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1743 | apply (subgoal_tac "b*q < b* (q' + 1) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1744 | apply (simp add: mult_less_cancel_left) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1745 | apply (subgoal_tac "b*q = r' - r + b'*q'") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1746 | prefer 2 apply simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1747 | apply (simp (no_asm_simp) add: right_distrib) | 
| 44766 | 1748 | apply (subst add_commute, rule add_less_le_mono, arith) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1749 | apply (rule mult_right_mono, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1750 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1751 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1752 | lemma zdiv_mono2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1753 | "[| (0::int) \<le> a; 0 < b'; b' \<le> b |] ==> a div b \<le> a div b'" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1754 | apply (subgoal_tac "b \<noteq> 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1755 | prefer 2 apply arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1756 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1757 | apply (cut_tac a = a and b = b' in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1758 | apply (rule zdiv_mono2_lemma) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1759 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1760 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1761 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1762 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1763 | lemma q_neg_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1764 | "[| b'*q' + r' < 0; 0 \<le> r'; 0 < b' |] ==> q' \<le> (0::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1765 | apply (subgoal_tac "b'*q' < 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1766 | apply (simp add: mult_less_0_iff, arith) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1767 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1768 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1769 | lemma zdiv_mono2_neg_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1770 | "[| b*q + r = b'*q' + r'; b'*q' + r' < 0; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1771 | r < b; 0 \<le> r'; 0 < b'; b' \<le> b |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1772 | ==> q' \<le> (q::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1773 | apply (frule q_neg_lemma, assumption+) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1774 | apply (subgoal_tac "b*q' < b* (q + 1) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1775 | apply (simp add: mult_less_cancel_left) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1776 | apply (simp add: right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1777 | apply (subgoal_tac "b*q' \<le> b'*q'") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1778 | prefer 2 apply (simp add: mult_right_mono_neg, arith) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1779 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1780 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1781 | lemma zdiv_mono2_neg: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1782 | "[| a < (0::int); 0 < b'; b' \<le> b |] ==> a div b' \<le> a div b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1783 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1784 | apply (cut_tac a = a and b = b' in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1785 | apply (rule zdiv_mono2_neg_lemma) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1786 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1787 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1788 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1789 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1790 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1791 | subsubsection {* More Algebraic Laws for div and mod *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1792 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1793 | text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1794 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1795 | lemma zmult1_lemma: | 
| 46552 | 1796 | "[| divmod_int_rel b c (q, r) |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1797 | ==> divmod_int_rel (a * b) c (a*q + a*r div c, a*r mod c)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1798 | by (auto simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib mult_ac) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1799 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1800 | lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1801 | apply (case_tac "c = 0", simp) | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1802 | apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN div_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1803 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1804 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1805 | text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1806 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1807 | lemma zadd1_lemma: | 
| 46552 | 1808 | "[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br) |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1809 | ==> divmod_int_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1810 | by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1811 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1812 | (*NOT suitable for rewriting: the RHS has an instance of the LHS*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1813 | lemma zdiv_zadd1_eq: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1814 | "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1815 | apply (case_tac "c = 0", simp) | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1816 | apply (blast intro: zadd1_lemma [OF divmod_int_rel_div_mod divmod_int_rel_div_mod] div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1817 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1818 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1819 | lemma posDivAlg_div_mod: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1820 | assumes "k \<ge> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1821 | and "l \<ge> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1822 | shows "posDivAlg k l = (k div l, k mod l)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1823 | proof (cases "l = 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1824 | case True then show ?thesis by (simp add: posDivAlg.simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1825 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1826 | case False with assms posDivAlg_correct | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1827 | have "divmod_int_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1828 | by simp | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1829 | from div_int_unique [OF this] mod_int_unique [OF this] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1830 | show ?thesis by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1831 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1832 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1833 | lemma negDivAlg_div_mod: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1834 | assumes "k < 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1835 | and "l > 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1836 | shows "negDivAlg k l = (k div l, k mod l)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1837 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1838 | from assms have "l \<noteq> 0" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1839 | from assms negDivAlg_correct | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1840 | have "divmod_int_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1841 | by simp | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1842 | from div_int_unique [OF this] mod_int_unique [OF this] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1843 | show ?thesis by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1844 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1845 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1846 | lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1847 | by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1848 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1849 | (* REVISIT: should this be generalized to all semiring_div types? *) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1850 | lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1851 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1852 | lemma zmod_zdiv_equality': | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1853 | "(m\<Colon>int) mod n = m - (m div n) * n" | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1854 | using mod_div_equality [of m n] by arith | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1855 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1856 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1857 | subsubsection {* Proving  @{term "a div (b*c) = (a div b) div c"} *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1858 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1859 | (*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1860 | 7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1861 | to cause particular problems.*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1862 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1863 | text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1864 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1865 | lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \<le> 0 |] ==> b*c < b*(q mod c) + r" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1866 | apply (subgoal_tac "b * (c - q mod c) < r * 1") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1867 | apply (simp add: algebra_simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1868 | apply (rule order_le_less_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1869 | apply (erule_tac [2] mult_strict_right_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1870 | apply (rule mult_left_mono_neg) | 
| 35216 | 1871 | using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1872 | apply (simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1873 | apply (simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1874 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1875 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1876 | lemma zmult2_lemma_aux2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1877 | "[| (0::int) < c; b < r; r \<le> 0 |] ==> b * (q mod c) + r \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1878 | apply (subgoal_tac "b * (q mod c) \<le> 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1879 | apply arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1880 | apply (simp add: mult_le_0_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1881 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1882 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1883 | lemma zmult2_lemma_aux3: "[| (0::int) < c; 0 \<le> r; r < b |] ==> 0 \<le> b * (q mod c) + r" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1884 | apply (subgoal_tac "0 \<le> b * (q mod c) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1885 | apply arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1886 | apply (simp add: zero_le_mult_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1887 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1888 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1889 | lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \<le> r; r < b |] ==> b * (q mod c) + r < b * c" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1890 | apply (subgoal_tac "r * 1 < b * (c - q mod c) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1891 | apply (simp add: right_diff_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1892 | apply (rule order_less_le_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1893 | apply (erule mult_strict_right_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1894 | apply (rule_tac [2] mult_left_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1895 | apply simp | 
| 35216 | 1896 | using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1897 | apply simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1898 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1899 | |
| 46552 | 1900 | lemma zmult2_lemma: "[| divmod_int_rel a b (q, r); 0 < c |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1901 | ==> divmod_int_rel a (b * c) (q div c, b*(q mod c) + r)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1902 | by (auto simp add: mult_ac divmod_int_rel_def linorder_neq_iff | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1903 | zero_less_mult_iff right_distrib [symmetric] | 
| 47139 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1904 | zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4 mult_less_0_iff split: split_if_asm) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1905 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1906 | lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1907 | apply (case_tac "b = 0", simp) | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1908 | apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN div_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1909 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1910 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1911 | lemma zmod_zmult2_eq: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1912 | "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1913 | apply (case_tac "b = 0", simp) | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1914 | apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN mod_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1915 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1916 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1917 | lemma div_pos_geq: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1918 | fixes k l :: int | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1919 | assumes "0 < l" and "l \<le> k" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1920 | shows "k div l = (k - l) div l + 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1921 | proof - | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1922 | have "k = (k - l) + l" by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1923 | then obtain j where k: "k = j + l" .. | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1924 | with assms show ?thesis by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1925 | qed | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1926 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1927 | lemma mod_pos_geq: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1928 | fixes k l :: int | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1929 | assumes "0 < l" and "l \<le> k" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1930 | shows "k mod l = (k - l) mod l" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1931 | proof - | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1932 | have "k = (k - l) + l" by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1933 | then obtain j where k: "k = j + l" .. | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1934 | with assms show ?thesis by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1935 | qed | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1936 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1937 | |
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 1938 | subsubsection {* Splitting Rules for div and mod *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1939 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1940 | text{*The proofs of the two lemmas below are essentially identical*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1941 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1942 | lemma split_pos_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1943 | "0<k ==> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1944 | P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i j)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1945 | apply (rule iffI, clarify) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1946 | apply (erule_tac P="P ?x ?y" in rev_mp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1947 | apply (subst mod_add_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1948 | apply (subst zdiv_zadd1_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1949 | apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1950 | txt{*converse direction*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1951 | apply (drule_tac x = "n div k" in spec) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1952 | apply (drule_tac x = "n mod k" in spec, simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1953 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1954 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1955 | lemma split_neg_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1956 | "k<0 ==> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1957 | P(n div k :: int)(n mod k) = (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i j)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1958 | apply (rule iffI, clarify) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1959 | apply (erule_tac P="P ?x ?y" in rev_mp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1960 | apply (subst mod_add_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1961 | apply (subst zdiv_zadd1_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1962 | apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1963 | txt{*converse direction*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1964 | apply (drule_tac x = "n div k" in spec) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1965 | apply (drule_tac x = "n mod k" in spec, simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1966 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1967 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1968 | lemma split_zdiv: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1969 | "P(n div k :: int) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1970 | ((k = 0 --> P 0) & | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1971 | (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i)) & | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1972 | (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1973 | apply (case_tac "k=0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1974 | apply (simp only: linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1975 | apply (erule disjE) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1976 | apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1977 | split_neg_lemma [of concl: "%x y. P x"]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1978 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1979 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1980 | lemma split_zmod: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1981 | "P(n mod k :: int) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1982 | ((k = 0 --> P n) & | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1983 | (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P j)) & | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1984 | (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P j)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1985 | apply (case_tac "k=0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1986 | apply (simp only: linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1987 | apply (erule disjE) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1988 | apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1989 | split_neg_lemma [of concl: "%x y. P y"]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1990 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1991 | |
| 33730 
1755ca4ec022
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 webertj parents: 
33728diff
changeset | 1992 | text {* Enable (lin)arith to deal with @{const div} and @{const mod}
 | 
| 
1755ca4ec022
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 webertj parents: 
33728diff
changeset | 1993 | when these are applied to some constant that is of the form | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1994 |   @{term "numeral k"}: *}
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1995 | declare split_zdiv [of _ _ "numeral k", arith_split] for k | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1996 | declare split_zmod [of _ _ "numeral k", arith_split] for k | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1997 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1998 | |
| 47166 | 1999 | subsubsection {* Computing @{text "div"} and @{text "mod"} with shifting *}
 | 
| 2000 | ||
| 2001 | lemma pos_divmod_int_rel_mult_2: | |
| 2002 | assumes "0 \<le> b" | |
| 2003 | assumes "divmod_int_rel a b (q, r)" | |
| 2004 | shows "divmod_int_rel (1 + 2*a) (2*b) (q, 1 + 2*r)" | |
| 2005 | using assms unfolding divmod_int_rel_def by auto | |
| 2006 | ||
| 2007 | lemma neg_divmod_int_rel_mult_2: | |
| 2008 | assumes "b \<le> 0" | |
| 2009 | assumes "divmod_int_rel (a + 1) b (q, r)" | |
| 2010 | shows "divmod_int_rel (1 + 2*a) (2*b) (q, 2*r - 1)" | |
| 2011 | using assms unfolding divmod_int_rel_def by auto | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2012 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2013 | text{*computing div by shifting *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2014 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2015 | lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a" | 
| 47166 | 2016 | using pos_divmod_int_rel_mult_2 [OF _ divmod_int_rel_div_mod] | 
| 2017 | by (rule div_int_unique) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2018 | |
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2019 | lemma neg_zdiv_mult_2: | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2020 | assumes A: "a \<le> (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a" | 
| 47166 | 2021 | using neg_divmod_int_rel_mult_2 [OF A divmod_int_rel_div_mod] | 
| 2022 | by (rule div_int_unique) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2023 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2024 | (* FIXME: add rules for negative numerals *) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2025 | lemma zdiv_numeral_Bit0 [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2026 | "numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2027 | numeral v div (numeral w :: int)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2028 | unfolding numeral.simps unfolding mult_2 [symmetric] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2029 | by (rule div_mult_mult1, simp) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2030 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2031 | lemma zdiv_numeral_Bit1 [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2032 | "numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2033 | (numeral v div (numeral w :: int))" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2034 | unfolding numeral.simps | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2035 | unfolding mult_2 [symmetric] add_commute [of _ 1] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2036 | by (rule pos_zdiv_mult_2, simp) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2037 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2038 | lemma pos_zmod_mult_2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2039 | fixes a b :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2040 | assumes "0 \<le> a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2041 | shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" | 
| 47166 | 2042 | using pos_divmod_int_rel_mult_2 [OF assms divmod_int_rel_div_mod] | 
| 2043 | by (rule mod_int_unique) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2044 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2045 | lemma neg_zmod_mult_2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2046 | fixes a b :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2047 | assumes "a \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2048 | shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" | 
| 47166 | 2049 | using neg_divmod_int_rel_mult_2 [OF assms divmod_int_rel_div_mod] | 
| 2050 | by (rule mod_int_unique) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2051 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2052 | (* FIXME: add rules for negative numerals *) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2053 | lemma zmod_numeral_Bit0 [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2054 | "numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2055 | (2::int) * (numeral v mod numeral w)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2056 | unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2057 | unfolding mult_2 [symmetric] by (rule mod_mult_mult1) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2058 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2059 | lemma zmod_numeral_Bit1 [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2060 | "numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2061 | 2 * (numeral v mod numeral w) + (1::int)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2062 | unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2063 | unfolding mult_2 [symmetric] add_commute [of _ 1] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2064 | by (rule pos_zmod_mult_2, simp) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2065 | |
| 39489 | 2066 | lemma zdiv_eq_0_iff: | 
| 2067 | "(i::int) div k = 0 \<longleftrightarrow> k=0 \<or> 0\<le>i \<and> i<k \<or> i\<le>0 \<and> k<i" (is "?L = ?R") | |
| 2068 | proof | |
| 2069 | assume ?L | |
| 2070 | have "?L \<longrightarrow> ?R" by (rule split_zdiv[THEN iffD2]) simp | |
| 2071 | with `?L` show ?R by blast | |
| 2072 | next | |
| 2073 | assume ?R thus ?L | |
| 2074 | by(auto simp: div_pos_pos_trivial div_neg_neg_trivial) | |
| 2075 | qed | |
| 2076 | ||
| 2077 | ||
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 2078 | subsubsection {* Quotients of Signs *}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2079 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2080 | lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2081 | apply (subgoal_tac "a div b \<le> -1", force) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2082 | apply (rule order_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2083 | apply (rule_tac a' = "-1" in zdiv_mono1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2084 | apply (auto simp add: div_eq_minus1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2085 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2086 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2087 | lemma div_nonneg_neg_le0: "[| (0::int) \<le> a; b < 0 |] ==> a div b \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2088 | by (drule zdiv_mono1_neg, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2089 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2090 | lemma div_nonpos_pos_le0: "[| (a::int) \<le> 0; b > 0 |] ==> a div b \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2091 | by (drule zdiv_mono1, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2092 | |
| 33804 | 2093 | text{* Now for some equivalences of the form @{text"a div b >=< 0 \<longleftrightarrow> \<dots>"}
 | 
| 2094 | conditional upon the sign of @{text a} or @{text b}. There are many more.
 | |
| 2095 | They should all be simp rules unless that causes too much search. *} | |
| 2096 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2097 | lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \<le> a div b) = (0 \<le> a)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2098 | apply auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2099 | apply (drule_tac [2] zdiv_mono1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2100 | apply (auto simp add: linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2101 | apply (simp (no_asm_use) add: linorder_not_less [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2102 | apply (blast intro: div_neg_pos_less0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2103 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2104 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2105 | lemma neg_imp_zdiv_nonneg_iff: | 
| 33804 | 2106 | "b < (0::int) ==> (0 \<le> a div b) = (a \<le> (0::int))" | 
| 47159 | 2107 | apply (subst div_minus_minus [symmetric]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2108 | apply (subst pos_imp_zdiv_nonneg_iff, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2109 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2110 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2111 | (*But not (a div b \<le> 0 iff a\<le>0); consider a=1, b=2 when a div b = 0.*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2112 | lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2113 | by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2114 | |
| 39489 | 2115 | lemma pos_imp_zdiv_pos_iff: | 
| 2116 | "0<k \<Longrightarrow> 0 < (i::int) div k \<longleftrightarrow> k \<le> i" | |
| 2117 | using pos_imp_zdiv_nonneg_iff[of k i] zdiv_eq_0_iff[of i k] | |
| 2118 | by arith | |
| 2119 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2120 | (*Again the law fails for \<le>: consider a = -1, b = -2 when a div b = 0*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2121 | lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2122 | by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2123 | |
| 33804 | 2124 | lemma nonneg1_imp_zdiv_pos_iff: | 
| 2125 | "(0::int) <= a \<Longrightarrow> (a div b > 0) = (a >= b & b>0)" | |
| 2126 | apply rule | |
| 2127 | apply rule | |
| 2128 | using div_pos_pos_trivial[of a b]apply arith | |
| 2129 | apply(cases "b=0")apply simp | |
| 2130 | using div_nonneg_neg_le0[of a b]apply arith | |
| 2131 | using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b]apply simp | |
| 2132 | done | |
| 2133 | ||
| 39489 | 2134 | lemma zmod_le_nonneg_dividend: "(m::int) \<ge> 0 ==> m mod k \<le> m" | 
| 2135 | apply (rule split_zmod[THEN iffD2]) | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44766diff
changeset | 2136 | apply(fastforce dest: q_pos_lemma intro: split_mult_pos_le) | 
| 39489 | 2137 | done | 
| 2138 | ||
| 2139 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2140 | subsubsection {* The Divides Relation *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2141 | |
| 47268 | 2142 | lemma dvd_neg_numeral_left [simp]: | 
| 2143 | fixes y :: "'a::comm_ring_1" | |
| 2144 | shows "(neg_numeral k) dvd y \<longleftrightarrow> (numeral k) dvd y" | |
| 2145 | unfolding neg_numeral_def minus_dvd_iff .. | |
| 2146 | ||
| 2147 | lemma dvd_neg_numeral_right [simp]: | |
| 2148 | fixes x :: "'a::comm_ring_1" | |
| 2149 | shows "x dvd (neg_numeral k) \<longleftrightarrow> x dvd (numeral k)" | |
| 2150 | unfolding neg_numeral_def dvd_minus_iff .. | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2151 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2152 | lemmas dvd_eq_mod_eq_0_numeral [simp] = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2153 | dvd_eq_mod_eq_0 [of "numeral x" "numeral y"] for x y | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2154 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2155 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2156 | subsubsection {* Further properties *}
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2157 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2158 | lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2159 | using zmod_zdiv_equality[where a="m" and b="n"] | 
| 47142 | 2160 | by (simp add: algebra_simps) (* FIXME: generalize *) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2161 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2162 | lemma zdiv_int: "int (a div b) = (int a) div (int b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2163 | apply (subst split_div, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2164 | apply (subst split_zdiv, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2165 | apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in unique_quotient) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2166 | apply (auto simp add: divmod_int_rel_def of_nat_mult) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2167 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2168 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2169 | lemma zmod_int: "int (a mod b) = (int a) mod (int b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2170 | apply (subst split_mod, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2171 | apply (subst split_zmod, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2172 | apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2173 | in unique_remainder) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2174 | apply (auto simp add: divmod_int_rel_def of_nat_mult) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2175 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2176 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2177 | lemma abs_div: "(y::int) dvd x \<Longrightarrow> abs (x div y) = abs x div abs y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2178 | by (unfold dvd_def, cases "y=0", auto simp add: abs_mult) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2179 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2180 | text{*Suggested by Matthias Daum*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2181 | lemma int_power_div_base: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2182 | "\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2183 | apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2184 | apply (erule ssubst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2185 | apply (simp only: power_add) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2186 | apply simp_all | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2187 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2188 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2189 | text {* by Brian Huffman *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2190 | lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2191 | by (rule mod_minus_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2192 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2193 | lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2194 | by (rule mod_diff_left_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2195 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2196 | lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2197 | by (rule mod_diff_right_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2198 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2199 | lemmas zmod_simps = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2200 | mod_add_left_eq [symmetric] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2201 | mod_add_right_eq [symmetric] | 
| 47142 | 2202 | mod_mult_right_eq[symmetric] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2203 | mod_mult_left_eq [symmetric] | 
| 47164 | 2204 | power_mod | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2205 | zminus_zmod zdiff_zmod_left zdiff_zmod_right | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2206 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2207 | text {* Distributive laws for function @{text nat}. *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2208 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2209 | lemma nat_div_distrib: "0 \<le> x \<Longrightarrow> nat (x div y) = nat x div nat y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2210 | apply (rule linorder_cases [of y 0]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2211 | apply (simp add: div_nonneg_neg_le0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2212 | apply simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2213 | apply (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2214 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2215 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2216 | (*Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2217 | lemma nat_mod_distrib: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2218 | "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> nat (x mod y) = nat x mod nat y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2219 | apply (case_tac "y = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2220 | apply (simp add: nat_eq_iff zmod_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2221 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2222 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2223 | text  {* transfer setup *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2224 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2225 | lemma transfer_nat_int_functions: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2226 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2227 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) mod (nat y) = nat (x mod y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2228 | by (auto simp add: nat_div_distrib nat_mod_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2229 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2230 | lemma transfer_nat_int_function_closures: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2231 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x div y >= 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2232 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x mod y >= 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2233 | apply (cases "y = 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2234 | apply (auto simp add: pos_imp_zdiv_nonneg_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2235 | apply (cases "y = 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2236 | apply auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2237 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2238 | |
| 35644 | 2239 | declare transfer_morphism_nat_int [transfer add return: | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2240 | transfer_nat_int_functions | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2241 | transfer_nat_int_function_closures | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2242 | ] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2243 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2244 | lemma transfer_int_nat_functions: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2245 | "(int x) div (int y) = int (x div y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2246 | "(int x) mod (int y) = int (x mod y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2247 | by (auto simp add: zdiv_int zmod_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2248 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2249 | lemma transfer_int_nat_function_closures: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2250 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x div y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2251 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x mod y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2252 | by (simp_all only: is_nat_def transfer_nat_int_function_closures) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2253 | |
| 35644 | 2254 | declare transfer_morphism_int_nat [transfer add return: | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2255 | transfer_int_nat_functions | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2256 | transfer_int_nat_function_closures | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2257 | ] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2258 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2259 | text{*Suggested by Matthias Daum*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2260 | lemma int_div_less_self: "\<lbrakk>0 < x; 1 < k\<rbrakk> \<Longrightarrow> x div k < (x::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2261 | apply (subgoal_tac "nat x div nat k < nat x") | 
| 34225 | 2262 | apply (simp add: nat_div_distrib [symmetric]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2263 | apply (rule Divides.div_less_dividend, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2264 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2265 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2266 | lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \<longleftrightarrow> n dvd x - y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2267 | proof | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2268 | assume H: "x mod n = y mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2269 | hence "x mod n - y mod n = 0" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2270 | hence "(x mod n - y mod n) mod n = 0" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2271 | hence "(x - y) mod n = 0" by (simp add: mod_diff_eq[symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2272 | thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2273 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2274 | assume H: "n dvd x - y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2275 | then obtain k where k: "x-y = n*k" unfolding dvd_def by blast | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2276 | hence "x = n*k + y" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2277 | hence "x mod n = (n*k + y) mod n" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2278 | thus "x mod n = y mod n" by (simp add: mod_add_left_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2279 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2280 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2281 | lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \<le> x" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2282 | shows "\<exists>q. x = y + n * q" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2283 | proof- | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2284 | from xy have th: "int x - int y = int (x - y)" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2285 | from xyn have "int x mod int n = int y mod int n" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 2286 | by (simp add: zmod_int [symmetric]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2287 | hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2288 | hence "n dvd x - y" by (simp add: th zdvd_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2289 | then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2290 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2291 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2292 | lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2293 | (is "?lhs = ?rhs") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2294 | proof | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2295 | assume H: "x mod n = y mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2296 |   {assume xy: "x \<le> y"
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2297 | from H have th: "y mod n = x mod n" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2298 | from nat_mod_eq_lemma[OF th xy] have ?rhs | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2299 | apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2300 | moreover | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2301 |   {assume xy: "y \<le> x"
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2302 | from nat_mod_eq_lemma[OF H xy] have ?rhs | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2303 | apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2304 | ultimately show ?rhs using linear[of x y] by blast | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2305 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2306 | assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2307 | hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2308 | thus ?lhs by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2309 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2310 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2311 | lemma div_nat_numeral [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2312 | "(numeral v :: nat) div numeral v' = nat (numeral v div numeral v')" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2313 | by (simp add: nat_div_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2314 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2315 | lemma one_div_nat_numeral [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2316 | "Suc 0 div numeral v' = nat (1 div numeral v')" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2317 | by (subst nat_div_distrib, simp_all) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2318 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2319 | lemma mod_nat_numeral [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2320 | "(numeral v :: nat) mod numeral v' = nat (numeral v mod numeral v')" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2321 | by (simp add: nat_mod_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2322 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2323 | lemma one_mod_nat_numeral [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2324 | "Suc 0 mod numeral v' = nat (1 mod numeral v')" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2325 | by (subst nat_mod_distrib) simp_all | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2326 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2327 | lemma mod_2_not_eq_zero_eq_one_int: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2328 | fixes k :: int | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2329 | shows "k mod 2 \<noteq> 0 \<longleftrightarrow> k mod 2 = 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2330 | by auto | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2331 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2332 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2333 | subsubsection {* Tools setup *}
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2334 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2335 | text {* Nitpick *}
 | 
| 34126 | 2336 | |
| 41792 
ff3cb0c418b7
renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
 blanchet parents: 
41550diff
changeset | 2337 | lemmas [nitpick_unfold] = dvd_eq_mod_eq_0 mod_div_equality' zmod_zdiv_equality' | 
| 34126 | 2338 | |
| 35673 | 2339 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2340 | subsubsection {* Code generation *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2341 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2342 | definition pdivmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2343 | "pdivmod k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2344 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2345 | lemma pdivmod_posDivAlg [code]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2346 | "pdivmod k l = (if l = 0 then (0, \<bar>k\<bar>) else posDivAlg \<bar>k\<bar> \<bar>l\<bar>)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2347 | by (subst posDivAlg_div_mod) (simp_all add: pdivmod_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2348 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2349 | lemma divmod_int_pdivmod: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2350 | apsnd ((op *) (sgn l)) (if 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2351 | then pdivmod k l | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2352 | else (let (r, s) = pdivmod k l in | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2353 | if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2354 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2355 | have aux: "\<And>q::int. - k = l * q \<longleftrightarrow> k = l * - q" by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2356 | show ?thesis | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2357 | by (simp add: divmod_int_mod_div pdivmod_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2358 | (auto simp add: aux not_less not_le zdiv_zminus1_eq_if | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2359 | zmod_zminus1_eq_if zdiv_zminus2_eq_if zmod_zminus2_eq_if) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2360 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2361 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2362 | lemma divmod_int_code [code]: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2363 | apsnd ((op *) (sgn l)) (if sgn k = sgn l | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2364 | then pdivmod k l | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2365 | else (let (r, s) = pdivmod k l in | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2366 | if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2367 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2368 | have "k \<noteq> 0 \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0 \<longleftrightarrow> sgn k = sgn l" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2369 | by (auto simp add: not_less sgn_if) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2370 | then show ?thesis by (simp add: divmod_int_pdivmod) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2371 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2372 | |
| 33364 | 2373 | code_modulename SML | 
| 2374 | Divides Arith | |
| 2375 | ||
| 2376 | code_modulename OCaml | |
| 2377 | Divides Arith | |
| 2378 | ||
| 2379 | code_modulename Haskell | |
| 2380 | Divides Arith | |
| 2381 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2382 | end |