| author | haftmann | 
| Fri, 21 Oct 2005 09:08:42 +0200 | |
| changeset 17943 | 48ec47217fe2 | 
| parent 17609 | 5156b731ebc8 | 
| child 18154 | 0c05abaf6244 | 
| permissions | -rw-r--r-- | 
| 3366 | 1 | (* Title: HOL/Divides.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 4 | Copyright 1999 University of Cambridge | 
| 3366 | 5 | |
| 6 | The division operators div, mod and the divides relation "dvd" | |
| 7 | *) | |
| 8 | ||
| 15131 | 9 | theory Divides | 
| 17508 | 10 | imports Datatype | 
| 15131 | 11 | begin | 
| 3366 | 12 | |
| 8902 | 13 | (*We use the same class for div and mod; | 
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 14 | moreover, dvd is defined whenever multiplication is*) | 
| 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 15 | axclass | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
10789diff
changeset | 16 | div < type | 
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 17 | |
| 13152 | 18 | instance nat :: div .. | 
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 19 | |
| 3366 | 20 | consts | 
| 13152 | 21 | div :: "'a::div \<Rightarrow> 'a \<Rightarrow> 'a" (infixl 70) | 
| 22 | mod :: "'a::div \<Rightarrow> 'a \<Rightarrow> 'a" (infixl 70) | |
| 23 | dvd :: "'a::times \<Rightarrow> 'a \<Rightarrow> bool" (infixl 50) | |
| 3366 | 24 | |
| 25 | ||
| 26 | defs | |
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 27 | |
| 13152 | 28 | mod_def: "m mod n == wfrec (trancl pred_nat) | 
| 7029 
08d4eb8500dd
new division laws taking advantage of (m div 0) = 0 and (m mod 0) = m
 paulson parents: 
6865diff
changeset | 29 | (%f j. if j<n | n=0 then j else f (j-n)) m" | 
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 30 | |
| 13152 | 31 | div_def: "m div n == wfrec (trancl pred_nat) | 
| 7029 
08d4eb8500dd
new division laws taking advantage of (m div 0) = 0 and (m mod 0) = m
 paulson parents: 
6865diff
changeset | 32 | (%f j. if j<n | n=0 then 0 else Suc (f (j-n))) m" | 
| 3366 | 33 | |
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 34 | (*The definition of dvd is polymorphic!*) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 35 | dvd_def: "m dvd n == \<exists>k. n = m*k" | 
| 3366 | 36 | |
| 10559 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 37 | (*This definition helps prove the harder properties of div and mod. | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 38 | It is copied from IntDiv.thy; should it be overloaded?*) | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 39 | constdefs | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 40 | quorem :: "(nat*nat) * (nat*nat) => bool" | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 41 | "quorem == %((a,b), (q,r)). | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 42 | a = b*q + r & | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 43 | (if 0<b then 0\<le>r & r<b else b<r & r \<le>0)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 44 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 45 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 46 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 47 | subsection{*Initial Lemmas*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 48 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 49 | lemmas wf_less_trans = | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 50 | def_wfrec [THEN trans, OF eq_reflection wf_pred_nat [THEN wf_trancl], | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 51 | standard] | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 52 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 53 | lemma mod_eq: "(%m. m mod n) = | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 54 | wfrec (trancl pred_nat) (%f j. if j<n | n=0 then j else f (j-n))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 55 | by (simp add: mod_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 56 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 57 | lemma div_eq: "(%m. m div n) = wfrec (trancl pred_nat) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 58 | (%f j. if j<n | n=0 then 0 else Suc (f (j-n)))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 59 | by (simp add: div_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 60 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 61 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 62 | (** Aribtrary definitions for division by zero. Useful to simplify | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 63 | certain equations **) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 64 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 65 | lemma DIVISION_BY_ZERO_DIV [simp]: "a div 0 = (0::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 66 | by (rule div_eq [THEN wf_less_trans], simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 67 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 68 | lemma DIVISION_BY_ZERO_MOD [simp]: "a mod 0 = (a::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 69 | by (rule mod_eq [THEN wf_less_trans], simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 70 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 71 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 72 | subsection{*Remainder*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 73 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 74 | lemma mod_less [simp]: "m<n ==> m mod n = (m::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 75 | by (rule mod_eq [THEN wf_less_trans], simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 76 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 77 | lemma mod_geq: "~ m < (n::nat) ==> m mod n = (m-n) mod n" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 78 | apply (case_tac "n=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 79 | apply (rule mod_eq [THEN wf_less_trans]) | 
| 15439 | 80 | apply (simp add: cut_apply less_eq) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 81 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 82 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 83 | (*Avoids the ugly ~m<n above*) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 84 | lemma le_mod_geq: "(n::nat) \<le> m ==> m mod n = (m-n) mod n" | 
| 16796 | 85 | by (simp add: mod_geq linorder_not_less) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 86 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 87 | lemma mod_if: "m mod (n::nat) = (if m<n then m else (m-n) mod n)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 88 | by (simp add: mod_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 89 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 90 | lemma mod_1 [simp]: "m mod Suc 0 = 0" | 
| 15251 | 91 | apply (induct "m") | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 92 | apply (simp_all (no_asm_simp) add: mod_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 93 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 94 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 95 | lemma mod_self [simp]: "n mod n = (0::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 96 | apply (case_tac "n=0") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 97 | apply (simp_all add: mod_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 98 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 99 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 100 | lemma mod_add_self2 [simp]: "(m+n) mod n = m mod (n::nat)" | 
| 15251 | 101 | apply (subgoal_tac "(n + m) mod n = (n+m-n) mod n") | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 102 | apply (simp add: add_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 103 | apply (subst mod_geq [symmetric], simp_all) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 104 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 105 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 106 | lemma mod_add_self1 [simp]: "(n+m) mod n = m mod (n::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 107 | by (simp add: add_commute mod_add_self2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 108 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 109 | lemma mod_mult_self1 [simp]: "(m + k*n) mod n = m mod (n::nat)" | 
| 15251 | 110 | apply (induct "k") | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 111 | apply (simp_all add: add_left_commute [of _ n]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 112 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 113 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 114 | lemma mod_mult_self2 [simp]: "(m + n*k) mod n = m mod (n::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 115 | by (simp add: mult_commute mod_mult_self1) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 116 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 117 | lemma mod_mult_distrib: "(m mod n) * (k::nat) = (m*k) mod (n*k)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 118 | apply (case_tac "n=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 119 | apply (case_tac "k=0", simp) | 
| 15251 | 120 | 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 | 121 | apply (subst mod_if, simp) | 
| 15439 | 122 | apply (simp add: mod_geq diff_mult_distrib) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 123 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 124 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 125 | lemma mod_mult_distrib2: "(k::nat) * (m mod n) = (k*m) mod (k*n)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 126 | by (simp add: mult_commute [of k] mod_mult_distrib) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 127 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 128 | lemma mod_mult_self_is_0 [simp]: "(m*n) mod n = (0::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 129 | apply (case_tac "n=0", simp) | 
| 15251 | 130 | apply (induct "m", simp) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 131 | apply (rename_tac "k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 132 | apply (cut_tac m = "k*n" and n = n in mod_add_self2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 133 | apply (simp add: add_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 134 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 135 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 136 | lemma mod_mult_self1_is_0 [simp]: "(n*m) mod n = (0::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 137 | by (simp add: mult_commute mod_mult_self_is_0) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 138 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 139 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 140 | subsection{*Quotient*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 141 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 142 | lemma div_less [simp]: "m<n ==> m div n = (0::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 143 | by (rule div_eq [THEN wf_less_trans], simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 144 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 145 | lemma div_geq: "[| 0<n; ~m<n |] ==> m div n = Suc((m-n) div n)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 146 | apply (rule div_eq [THEN wf_less_trans]) | 
| 15439 | 147 | apply (simp add: cut_apply less_eq) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 148 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 149 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 150 | (*Avoids the ugly ~m<n above*) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 151 | lemma le_div_geq: "[| 0<n; n\<le>m |] ==> m div n = Suc((m-n) div n)" | 
| 16796 | 152 | by (simp add: div_geq linorder_not_less) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 153 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 154 | lemma div_if: "0<n ==> m div n = (if m<n then 0 else Suc((m-n) div n))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 155 | by (simp add: div_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 156 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 157 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 158 | (*Main Result about quotient and remainder.*) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 159 | lemma mod_div_equality: "(m div n)*n + m mod n = (m::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 160 | apply (case_tac "n=0", simp) | 
| 15251 | 161 | 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 | 162 | apply (subst mod_if) | 
| 15439 | 163 | apply (simp_all (no_asm_simp) add: add_assoc div_geq add_diff_inverse) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 164 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 165 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 166 | lemma mod_div_equality2: "n * (m div n) + m mod n = (m::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 167 | apply(cut_tac m = m and n = n in mod_div_equality) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 168 | apply(simp add: mult_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 169 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 170 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 171 | subsection{*Simproc for Cancelling Div and Mod*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 172 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 173 | lemma div_mod_equality: "((m div n)*n + m mod n) + k = (m::nat) + k" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 174 | apply(simp add: mod_div_equality) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 175 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 176 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 177 | lemma div_mod_equality2: "(n*(m div n) + m mod n) + k = (m::nat) + k" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 178 | apply(simp add: mod_div_equality2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 179 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 180 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 181 | ML | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 182 | {*
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 183 | val div_mod_equality = thm "div_mod_equality"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 184 | val div_mod_equality2 = thm "div_mod_equality2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 185 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 186 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 187 | structure CancelDivModData = | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 188 | struct | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 189 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 190 | val div_name = "Divides.op div"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 191 | val mod_name = "Divides.op mod"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 192 | val mk_binop = HOLogic.mk_binop; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 193 | val mk_sum = NatArithUtils.mk_sum; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 194 | val dest_sum = NatArithUtils.dest_sum; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 195 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 196 | (*logic*) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 197 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 198 | val div_mod_eqs = map mk_meta_eq [div_mod_equality,div_mod_equality2] | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 199 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 200 | val trans = trans | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 201 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 202 | val prove_eq_sums = | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 203 | let val simps = add_0 :: add_0_right :: add_ac | 
| 17609 
5156b731ebc8
Provers/cancel_sums.ML: Simplifier.inherit_bounds;
 wenzelm parents: 
17508diff
changeset | 204 | in NatArithUtils.prove_conv all_tac (NatArithUtils.simp_all_tac simps) end; | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 205 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 206 | end; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 207 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 208 | structure CancelDivMod = CancelDivModFun(CancelDivModData); | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 209 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 210 | val cancel_div_mod_proc = NatArithUtils.prep_simproc | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 211 |       ("cancel_div_mod", ["(m::nat) + n"], CancelDivMod.proc);
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 212 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 213 | Addsimprocs[cancel_div_mod_proc]; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 214 | *} | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 215 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 216 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 217 | (* a simple rearrangement of mod_div_equality: *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 218 | lemma mult_div_cancel: "(n::nat) * (m div n) = m - (m mod n)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 219 | by (cut_tac m = m and n = n in mod_div_equality2, arith) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 220 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 221 | lemma mod_less_divisor [simp]: "0<n ==> m mod n < (n::nat)" | 
| 15251 | 222 | 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 | 223 | apply (case_tac "na<n", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 224 | txt{*case @{term "n \<le> na"}*}
 | 
| 15439 | 225 | apply (simp add: mod_geq) | 
| 226 | done | |
| 227 | ||
| 228 | lemma mod_le_divisor[simp]: "0 < n \<Longrightarrow> m mod n \<le> (n::nat)" | |
| 229 | apply(drule mod_less_divisor[where m = m]) | |
| 230 | apply simp | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 231 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 232 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 233 | lemma div_mult_self_is_m [simp]: "0<n ==> (m*n) div n = (m::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 234 | by (cut_tac m = "m*n" and n = n in mod_div_equality, auto) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 235 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 236 | lemma div_mult_self1_is_m [simp]: "0<n ==> (n*m) div n = (m::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 237 | by (simp add: mult_commute div_mult_self_is_m) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 238 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 239 | (*mod_mult_distrib2 above is the counterpart for remainder*) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 240 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 241 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 242 | subsection{*Proving facts about Quotient and Remainder*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 243 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 244 | lemma unique_quotient_lemma: | 
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15439diff
changeset | 245 | "[| b*q' + r' \<le> b*q + r; x < b; r < b |] | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 246 | ==> q' \<le> (q::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 247 | apply (rule leI) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 248 | apply (subst less_iff_Suc_add) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 249 | apply (auto simp add: add_mult_distrib2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 250 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 251 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 252 | lemma unique_quotient: | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 253 | "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); 0 < b |] | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 254 | ==> q = q'" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 255 | apply (simp add: split_ifs quorem_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 256 | apply (blast intro: order_antisym | 
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15439diff
changeset | 257 | dest: order_eq_refl [THEN unique_quotient_lemma] sym) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 258 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 259 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 260 | lemma unique_remainder: | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 261 | "[| quorem ((a,b), (q,r)); quorem ((a,b), (q',r')); 0 < b |] | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 262 | ==> r = r'" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 263 | apply (subgoal_tac "q = q'") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 264 | prefer 2 apply (blast intro: unique_quotient) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 265 | apply (simp add: quorem_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 266 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 267 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 268 | lemma quorem_div_mod: "0 < b ==> quorem ((a, b), (a div b, a mod b))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 269 | by (auto simp add: quorem_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 270 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 271 | lemma quorem_div: "[| quorem((a,b),(q,r)); 0 < b |] ==> a div b = q" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 272 | by (simp add: quorem_div_mod [THEN unique_quotient]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 273 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 274 | lemma quorem_mod: "[| quorem((a,b),(q,r)); 0 < b |] ==> a mod b = r" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 275 | by (simp add: quorem_div_mod [THEN unique_remainder]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 276 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 277 | (** A dividend of zero **) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 278 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 279 | lemma div_0 [simp]: "0 div m = (0::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 280 | by (case_tac "m=0", simp_all) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 281 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 282 | lemma mod_0 [simp]: "0 mod m = (0::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 283 | by (case_tac "m=0", simp_all) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 284 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 285 | (** proving (a*b) div c = a * (b div c) + a * (b mod c) **) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 286 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 287 | lemma quorem_mult1_eq: | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 288 | "[| quorem((b,c),(q,r)); 0 < c |] | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 289 | ==> quorem ((a*b, c), (a*q + a*r div c, a*r mod c))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 290 | apply (auto simp add: split_ifs mult_ac quorem_def add_mult_distrib2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 291 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 292 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 293 | lemma div_mult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 294 | apply (case_tac "c = 0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 295 | apply (blast intro: quorem_div_mod [THEN quorem_mult1_eq, THEN quorem_div]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 296 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 297 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 298 | lemma mod_mult1_eq: "(a*b) mod c = a*(b mod c) mod (c::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 299 | apply (case_tac "c = 0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 300 | apply (blast intro: quorem_div_mod [THEN quorem_mult1_eq, THEN quorem_mod]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 301 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 302 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 303 | lemma mod_mult1_eq': "(a*b) mod (c::nat) = ((a mod c) * b) mod c" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 304 | apply (rule trans) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 305 | apply (rule_tac s = "b*a mod c" in trans) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 306 | apply (rule_tac [2] mod_mult1_eq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 307 | apply (simp_all (no_asm) add: mult_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 308 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 309 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 310 | lemma mod_mult_distrib_mod: "(a*b) mod (c::nat) = ((a mod c) * (b mod c)) mod c" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 311 | apply (rule mod_mult1_eq' [THEN trans]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 312 | apply (rule mod_mult1_eq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 313 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 314 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 315 | (** proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) **) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 316 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 317 | lemma quorem_add1_eq: | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 318 | "[| quorem((a,c),(aq,ar)); quorem((b,c),(bq,br)); 0 < c |] | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 319 | ==> quorem ((a+b, c), (aq + bq + (ar+br) div c, (ar+br) mod c))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 320 | by (auto simp add: split_ifs mult_ac quorem_def add_mult_distrib2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 321 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 322 | (*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 | 323 | lemma div_add1_eq: | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 324 | "(a+b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 325 | apply (case_tac "c = 0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 326 | apply (blast intro: quorem_add1_eq [THEN quorem_div] quorem_div_mod quorem_div_mod) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 327 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 328 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 329 | lemma mod_add1_eq: "(a+b) mod (c::nat) = (a mod c + b mod c) mod c" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 330 | apply (case_tac "c = 0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 331 | apply (blast intro: quorem_div_mod quorem_div_mod | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 332 | quorem_add1_eq [THEN quorem_mod]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 333 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 334 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 335 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 336 | subsection{*Proving @{term "a div (b*c) = (a div b) div c"}*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 337 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 338 | (** first, a lemma to bound the remainder **) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 339 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 340 | lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 341 | apply (cut_tac m = q and n = c in mod_less_divisor) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 342 | apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 343 | apply (erule_tac P = "%x. ?lhs < ?rhs x" in ssubst) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 344 | apply (simp add: add_mult_distrib2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 345 | done | 
| 10559 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 346 | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 347 | lemma quorem_mult2_eq: "[| quorem ((a,b), (q,r)); 0 < b; 0 < c |] | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 348 | ==> quorem ((a, b*c), (q div c, b*(q mod c) + r))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 349 | apply (auto simp add: mult_ac quorem_def add_mult_distrib2 [symmetric] mod_lemma) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 350 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 351 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 352 | lemma div_mult2_eq: "a div (b*c) = (a div b) div (c::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 353 | apply (case_tac "b=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 354 | apply (case_tac "c=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 355 | apply (force simp add: quorem_div_mod [THEN quorem_mult2_eq, THEN quorem_div]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 356 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 357 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 358 | lemma mod_mult2_eq: "a mod (b*c) = b*(a div b mod c) + a mod (b::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 359 | apply (case_tac "b=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 360 | apply (case_tac "c=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 361 | apply (auto simp add: mult_commute quorem_div_mod [THEN quorem_mult2_eq, THEN quorem_mod]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 362 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 363 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 364 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 365 | subsection{*Cancellation of Common Factors in Division*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 366 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 367 | lemma div_mult_mult_lemma: | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 368 | "[| (0::nat) < b; 0 < c |] ==> (c*a) div (c*b) = a div b" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 369 | by (auto simp add: div_mult2_eq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 370 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 371 | lemma div_mult_mult1 [simp]: "(0::nat) < c ==> (c*a) div (c*b) = a div b" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 372 | apply (case_tac "b = 0") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 373 | apply (auto simp add: linorder_neq_iff [of b] div_mult_mult_lemma) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 374 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 375 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 376 | lemma div_mult_mult2 [simp]: "(0::nat) < c ==> (a*c) div (b*c) = a div b" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 377 | apply (drule div_mult_mult1) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 378 | apply (auto simp add: mult_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 379 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 380 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 381 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 382 | (*Distribution of Factors over Remainders: | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 383 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 384 | Could prove these as in Integ/IntDiv.ML, but we already have | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 385 | mod_mult_distrib and mod_mult_distrib2 above! | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 386 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 387 | Goal "(c*a) mod (c*b) = (c::nat) * (a mod b)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 388 | qed "mod_mult_mult1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 389 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 390 | Goal "(a*c) mod (b*c) = (a mod b) * (c::nat)"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 391 | qed "mod_mult_mult2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 392 | ***) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 393 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 394 | subsection{*Further Facts about Quotient and Remainder*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 395 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 396 | lemma div_1 [simp]: "m div Suc 0 = m" | 
| 15251 | 397 | apply (induct "m") | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 398 | apply (simp_all (no_asm_simp) add: div_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 399 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 400 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 401 | lemma div_self [simp]: "0<n ==> n div n = (1::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 402 | by (simp add: div_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 403 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 404 | lemma div_add_self2: "0<n ==> (m+n) div n = Suc (m div n)" | 
| 15251 | 405 | apply (subgoal_tac "(n + m) div n = Suc ((n+m-n) div n) ") | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 406 | apply (simp add: add_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 407 | apply (subst div_geq [symmetric], simp_all) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 408 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 409 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 410 | lemma div_add_self1: "0<n ==> (n+m) div n = Suc (m div n)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 411 | by (simp add: add_commute div_add_self2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 412 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 413 | lemma div_mult_self1 [simp]: "!!n::nat. 0<n ==> (m + k*n) div n = k + m div n" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 414 | apply (subst div_add1_eq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 415 | apply (subst div_mult1_eq, simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 416 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 417 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 418 | lemma div_mult_self2 [simp]: "0<n ==> (m + n*k) div n = k + m div (n::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 419 | by (simp add: mult_commute div_mult_self1) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 420 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 421 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 422 | (* Monotonicity of div in first argument *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 423 | lemma div_le_mono [rule_format (no_asm)]: | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 424 | "\<forall>m::nat. m \<le> n --> (m div k) \<le> (n div k)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 425 | apply (case_tac "k=0", simp) | 
| 15251 | 426 | 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 | 427 | apply (case_tac "n<k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 428 | (* 1 case n<k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 429 | apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 430 | (* 2 case n >= k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 431 | apply (case_tac "m<k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 432 | (* 2.1 case m<k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 433 | apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 434 | (* 2.2 case m>=k *) | 
| 15439 | 435 | 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 | 436 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 437 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 438 | (* Antimonotonicity of div in second argument *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 439 | 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 | 440 | apply (subgoal_tac "0<n") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 441 | prefer 2 apply simp | 
| 15251 | 442 | 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 | 443 | apply (rename_tac "k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 444 | apply (case_tac "k<n", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 445 | apply (subgoal_tac "~ (k<m) ") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 446 | prefer 2 apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 447 | apply (simp add: div_geq) | 
| 15251 | 448 | 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 | 449 | prefer 2 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 450 | 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 | 451 | apply (rule le_trans, simp) | 
| 15439 | 452 | apply (simp) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 453 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 454 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 455 | 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 | 456 | apply (case_tac "n=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 457 | 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 | 458 | apply (rule div_le_mono2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 459 | apply (simp_all (no_asm_simp)) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 460 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 461 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 462 | (* Similar for "less than" *) | 
| 17085 | 463 | lemma div_less_dividend [rule_format]: | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 464 | "!!n::nat. 1<n ==> 0 < m --> m div n < m" | 
| 15251 | 465 | apply (induct_tac m rule: nat_less_induct) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 466 | apply (rename_tac "m") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 467 | apply (case_tac "m<n", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 468 | apply (subgoal_tac "0<n") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 469 | prefer 2 apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 470 | apply (simp add: div_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 471 | apply (case_tac "n<m") | 
| 15251 | 472 | 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 | 473 | apply (rule impI less_trans_Suc)+ | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 474 | apply assumption | 
| 15439 | 475 | apply (simp_all) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 476 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 477 | |
| 17085 | 478 | declare div_less_dividend [simp] | 
| 479 | ||
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 480 | 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 | 481 | 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 | 482 | apply (case_tac "n=0", simp) | 
| 15251 | 483 | 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 | 484 | apply (case_tac "Suc (na) <n") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 485 | (* case Suc(na) < n *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 486 | 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 | 487 | (* case n \<le> Suc(na) *) | 
| 16796 | 488 | apply (simp add: linorder_not_less le_Suc_eq mod_geq) | 
| 15439 | 489 | 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 | 490 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 491 | |
| 14437 | 492 | lemma nat_mod_div_trivial [simp]: "m mod n div n = (0 :: nat)" | 
| 493 | by (case_tac "n=0", auto) | |
| 494 | ||
| 495 | lemma nat_mod_mod_trivial [simp]: "m mod n mod n = (m mod n :: nat)" | |
| 496 | by (case_tac "n=0", auto) | |
| 497 | ||
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 498 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 499 | subsection{*The Divides Relation*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 500 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 501 | lemma dvdI [intro?]: "n = m * k ==> m dvd n" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 502 | by (unfold dvd_def, blast) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 503 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 504 | lemma dvdE [elim?]: "!!P. [|m dvd n; !!k. n = m*k ==> P|] ==> P" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 505 | by (unfold dvd_def, blast) | 
| 13152 | 506 | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 507 | lemma dvd_0_right [iff]: "m dvd (0::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 508 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 509 | apply (blast intro: mult_0_right [symmetric]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 510 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 511 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 512 | lemma dvd_0_left: "0 dvd m ==> m = (0::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 513 | by (force simp add: dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 514 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 515 | lemma dvd_0_left_iff [iff]: "(0 dvd (m::nat)) = (m = 0)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 516 | by (blast intro: dvd_0_left) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 517 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 518 | lemma dvd_1_left [iff]: "Suc 0 dvd k" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 519 | by (unfold dvd_def, simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 520 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 521 | lemma dvd_1_iff_1 [simp]: "(m dvd Suc 0) = (m = Suc 0)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 522 | by (simp add: dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 523 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 524 | lemma dvd_refl [simp]: "m dvd (m::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 525 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 526 | apply (blast intro: mult_1_right [symmetric]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 527 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 528 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 529 | lemma dvd_trans [trans]: "[| m dvd n; n dvd p |] ==> m dvd (p::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 530 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 531 | apply (blast intro: mult_assoc) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 532 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 533 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 534 | lemma dvd_anti_sym: "[| m dvd n; n dvd m |] ==> m = (n::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 535 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 536 | apply (force dest: mult_eq_self_implies_10 simp add: mult_assoc mult_eq_1_iff) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 537 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 538 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 539 | lemma dvd_add: "[| k dvd m; k dvd n |] ==> k dvd (m+n :: nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 540 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 541 | apply (blast intro: add_mult_distrib2 [symmetric]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 542 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 543 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 544 | lemma dvd_diff: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 545 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 546 | apply (blast intro: diff_mult_distrib2 [symmetric]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 547 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 548 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 549 | lemma dvd_diffD: "[| k dvd m-n; k dvd n; n\<le>m |] ==> k dvd (m::nat)" | 
| 16796 | 550 | apply (erule linorder_not_less [THEN iffD2, THEN add_diff_inverse, THEN subst]) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 551 | apply (blast intro: dvd_add) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 552 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 553 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 554 | lemma dvd_diffD1: "[| k dvd m-n; k dvd m; n\<le>m |] ==> k dvd (n::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 555 | by (drule_tac m = m in dvd_diff, auto) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 556 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 557 | lemma dvd_mult: "k dvd n ==> k dvd (m*n :: nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 558 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 559 | apply (blast intro: mult_left_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 560 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 561 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 562 | lemma dvd_mult2: "k dvd m ==> k dvd (m*n :: nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 563 | apply (subst mult_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 564 | apply (erule dvd_mult) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 565 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 566 | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16796diff
changeset | 567 | lemma dvd_triv_right [iff]: "k dvd (m*k :: nat)" | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16796diff
changeset | 568 | by (rule dvd_refl [THEN dvd_mult]) | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16796diff
changeset | 569 | |
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16796diff
changeset | 570 | lemma dvd_triv_left [iff]: "k dvd (k*m :: nat)" | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16796diff
changeset | 571 | by (rule dvd_refl [THEN dvd_mult2]) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 572 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 573 | lemma dvd_reduce: "(k dvd n + k) = (k dvd (n::nat))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 574 | apply (rule iffI) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 575 | apply (erule_tac [2] dvd_add) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 576 | apply (rule_tac [2] dvd_refl) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 577 | apply (subgoal_tac "n = (n+k) -k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 578 | prefer 2 apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 579 | apply (erule ssubst) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 580 | apply (erule dvd_diff) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 581 | apply (rule dvd_refl) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 582 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 583 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 584 | lemma dvd_mod: "!!n::nat. [| f dvd m; f dvd n |] ==> f dvd m mod n" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 585 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 586 | apply (case_tac "n=0", auto) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 587 | apply (blast intro: mod_mult_distrib2 [symmetric]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 588 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 589 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 590 | lemma dvd_mod_imp_dvd: "[| (k::nat) dvd m mod n; k dvd n |] ==> k dvd m" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 591 | apply (subgoal_tac "k dvd (m div n) *n + m mod n") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 592 | apply (simp add: mod_div_equality) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 593 | apply (simp only: dvd_add dvd_mult) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 594 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 595 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 596 | lemma dvd_mod_iff: "k dvd n ==> ((k::nat) dvd m mod n) = (k dvd m)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 597 | by (blast intro: dvd_mod_imp_dvd dvd_mod) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 598 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 599 | lemma dvd_mult_cancel: "!!k::nat. [| k*m dvd k*n; 0<k |] ==> m dvd n" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 600 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 601 | apply (erule exE) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 602 | apply (simp add: mult_ac) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 603 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 604 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 605 | lemma dvd_mult_cancel1: "0<m ==> (m*n dvd m) = (n = (1::nat))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 606 | apply auto | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 607 | apply (subgoal_tac "m*n dvd m*1") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 608 | apply (drule dvd_mult_cancel, auto) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 609 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 610 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 611 | lemma dvd_mult_cancel2: "0<m ==> (n*m dvd m) = (n = (1::nat))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 612 | apply (subst mult_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 613 | apply (erule dvd_mult_cancel1) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 614 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 615 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 616 | lemma mult_dvd_mono: "[| i dvd m; j dvd n|] ==> i*j dvd (m*n :: nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 617 | apply (unfold dvd_def, clarify) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 618 | apply (rule_tac x = "k*ka" in exI) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 619 | apply (simp add: mult_ac) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 620 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 621 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 622 | lemma dvd_mult_left: "(i*j :: nat) dvd k ==> i dvd k" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 623 | by (simp add: dvd_def mult_assoc, blast) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 624 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 625 | lemma dvd_mult_right: "(i*j :: nat) dvd k ==> j dvd k" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 626 | apply (unfold dvd_def, clarify) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 627 | apply (rule_tac x = "i*k" in exI) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 628 | apply (simp add: mult_ac) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 629 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 630 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 631 | lemma dvd_imp_le: "[| k dvd n; 0 < n |] ==> k \<le> (n::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 632 | apply (unfold dvd_def, clarify) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 633 | apply (simp_all (no_asm_use) add: zero_less_mult_iff) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 634 | apply (erule conjE) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 635 | apply (rule le_trans) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 636 | apply (rule_tac [2] le_refl [THEN mult_le_mono]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 637 | apply (erule_tac [2] Suc_leI, simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 638 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 639 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 640 | lemma dvd_eq_mod_eq_0: "!!k::nat. (k dvd n) = (n mod k = 0)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 641 | apply (unfold dvd_def) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 642 | apply (case_tac "k=0", simp, safe) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 643 | apply (simp add: mult_commute) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 644 | apply (rule_tac t = n and n1 = k in mod_div_equality [THEN subst]) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 645 | apply (subst mult_commute, simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 646 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 647 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 648 | lemma dvd_mult_div_cancel: "n dvd m ==> n * (m div n) = (m::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 649 | apply (subgoal_tac "m mod n = 0") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 650 | apply (simp add: mult_div_cancel) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 651 | apply (simp only: dvd_eq_mod_eq_0) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 652 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 653 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 654 | lemma mod_eq_0_iff: "(m mod d = 0) = (\<exists>q::nat. m = d*q)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 655 | 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 | 656 | |
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16796diff
changeset | 657 | lemmas mod_eq_0D = mod_eq_0_iff [THEN iffD1] | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16796diff
changeset | 658 | declare mod_eq_0D [dest!] | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 659 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 660 | (*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 | 661 | lemma mod_eqD: "(m mod d = r) ==> \<exists>q::nat. m = r + q*d" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 662 | apply (cut_tac m = m in mod_div_equality) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 663 | apply (simp only: add_ac) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 664 | apply (blast intro: sym) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 665 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 666 | |
| 14131 | 667 | |
| 13152 | 668 | lemma split_div: | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 669 | "P(n div k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 670 | ((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 | 671 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 672 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 673 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 674 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 675 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 676 | assume "k = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 677 | with P show ?Q by(simp add:DIVISION_BY_ZERO_DIV) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 678 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 679 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 680 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 681 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 682 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 683 | 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 | 684 | show "P i" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 685 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 686 | assume "i = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 687 | with n j P show "P i" by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 688 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 689 | assume "i \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 690 | with not0 n j P show "P i" by(simp add:add_ac) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 691 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 692 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 693 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 694 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 695 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 696 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 697 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 698 | assume "k = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 699 | with Q show ?P by(simp add:DIVISION_BY_ZERO_DIV) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 700 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 701 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 702 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 703 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 704 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 705 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 706 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 707 | |
| 13882 | 708 | lemma split_div_lemma: | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 709 | "0 < n \<Longrightarrow> (n * q \<le> m \<and> m < n * (Suc q)) = (q = ((m::nat) div n))" | 
| 13882 | 710 | apply (rule iffI) | 
| 711 | apply (rule_tac a=m and r = "m - n * q" and r' = "m mod n" in unique_quotient) | |
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15439diff
changeset | 712 | prefer 3; apply assumption | 
| 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15439diff
changeset | 713 | apply (simp_all add: quorem_def) | 
| 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15439diff
changeset | 714 | apply arith | 
| 13882 | 715 | apply (rule conjI) | 
| 716 | apply (rule_tac P="%x. n * (m div n) \<le> x" in | |
| 717 | subst [OF mod_div_equality [of _ n]]) | |
| 718 | apply (simp only: add: mult_ac) | |
| 719 | apply (rule_tac P="%x. x < n + n * (m div n)" in | |
| 720 | subst [OF mod_div_equality [of _ n]]) | |
| 721 | apply (simp only: add: mult_ac add_ac) | |
| 14208 | 722 | apply (rule add_less_mono1, simp) | 
| 13882 | 723 | done | 
| 724 | ||
| 725 | theorem split_div': | |
| 726 | "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 | 727 | (\<exists>q. (n * q \<le> m \<and> m < n * (Suc q)) \<and> P q))" | 
| 13882 | 728 | apply (case_tac "0 < n") | 
| 729 | apply (simp only: add: split_div_lemma) | |
| 730 | apply (simp_all add: DIVISION_BY_ZERO_DIV) | |
| 731 | done | |
| 732 | ||
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 733 | lemma split_mod: | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 734 | "P(n mod k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 735 | ((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 | 736 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 737 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 738 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 739 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 740 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 741 | assume "k = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 742 | with P show ?Q by(simp add:DIVISION_BY_ZERO_MOD) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 743 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 744 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 745 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 746 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 747 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 748 | assume "n = k*i + j" "j < k" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 749 | 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 | 750 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 751 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 752 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 753 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 754 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 755 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 756 | assume "k = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 757 | with Q show ?P by(simp add:DIVISION_BY_ZERO_MOD) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 758 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 759 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 760 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 761 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 762 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 763 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 764 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 765 | |
| 13882 | 766 | theorem mod_div_equality': "(m::nat) mod n = m - (m div n) * n" | 
| 767 | apply (rule_tac P="%x. m mod n = x - (m div n) * n" in | |
| 768 | subst [OF mod_div_equality [of _ n]]) | |
| 769 | apply arith | |
| 770 | done | |
| 771 | ||
| 14640 | 772 | subsection {*An ``induction'' law for modulus arithmetic.*}
 | 
| 773 | ||
| 774 | lemma mod_induct_0: | |
| 775 | assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)" | |
| 776 | and base: "P i" and i: "i<p" | |
| 777 | shows "P 0" | |
| 778 | proof (rule ccontr) | |
| 779 | assume contra: "\<not>(P 0)" | |
| 780 | from i have p: "0<p" by simp | |
| 781 | have "\<forall>k. 0<k \<longrightarrow> \<not> P (p-k)" (is "\<forall>k. ?A k") | |
| 782 | proof | |
| 783 | fix k | |
| 784 | show "?A k" | |
| 785 | proof (induct k) | |
| 786 | show "?A 0" by simp -- "by contradiction" | |
| 787 | next | |
| 788 | fix n | |
| 789 | assume ih: "?A n" | |
| 790 | show "?A (Suc n)" | |
| 791 | proof (clarsimp) | |
| 792 | assume y: "P (p - Suc n)" | |
| 793 | have n: "Suc n < p" | |
| 794 | proof (rule ccontr) | |
| 795 | assume "\<not>(Suc n < p)" | |
| 796 | hence "p - Suc n = 0" | |
| 797 | by simp | |
| 798 | with y contra show "False" | |
| 799 | by simp | |
| 800 | qed | |
| 801 | hence n2: "Suc (p - Suc n) = p-n" by arith | |
| 802 | from p have "p - Suc n < p" by arith | |
| 803 | with y step have z: "P ((Suc (p - Suc n)) mod p)" | |
| 804 | by blast | |
| 805 | show "False" | |
| 806 | proof (cases "n=0") | |
| 807 | case True | |
| 808 | with z n2 contra show ?thesis by simp | |
| 809 | next | |
| 810 | case False | |
| 811 | with p have "p-n < p" by arith | |
| 812 | with z n2 False ih show ?thesis by simp | |
| 813 | qed | |
| 814 | qed | |
| 815 | qed | |
| 816 | qed | |
| 817 | moreover | |
| 818 | from i obtain k where "0<k \<and> i+k=p" | |
| 819 | by (blast dest: less_imp_add_positive) | |
| 820 | hence "0<k \<and> i=p-k" by auto | |
| 821 | moreover | |
| 822 | note base | |
| 823 | ultimately | |
| 824 | show "False" by blast | |
| 825 | qed | |
| 826 | ||
| 827 | lemma mod_induct: | |
| 828 | assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)" | |
| 829 | and base: "P i" and i: "i<p" and j: "j<p" | |
| 830 | shows "P j" | |
| 831 | proof - | |
| 832 | have "\<forall>j<p. P j" | |
| 833 | proof | |
| 834 | fix j | |
| 835 | show "j<p \<longrightarrow> P j" (is "?A j") | |
| 836 | proof (induct j) | |
| 837 | from step base i show "?A 0" | |
| 838 | by (auto elim: mod_induct_0) | |
| 839 | next | |
| 840 | fix k | |
| 841 | assume ih: "?A k" | |
| 842 | show "?A (Suc k)" | |
| 843 | proof | |
| 844 | assume suc: "Suc k < p" | |
| 845 | hence k: "k<p" by simp | |
| 846 | with ih have "P k" .. | |
| 847 | with step k have "P (Suc k mod p)" | |
| 848 | by blast | |
| 849 | moreover | |
| 850 | from suc have "Suc k mod p = Suc k" | |
| 851 | by simp | |
| 852 | ultimately | |
| 853 | show "P (Suc k)" by simp | |
| 854 | qed | |
| 855 | qed | |
| 856 | qed | |
| 857 | with j show ?thesis by blast | |
| 858 | qed | |
| 859 | ||
| 860 | ||
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 861 | ML | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 862 | {*
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 863 | val div_def = thm "div_def" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 864 | val mod_def = thm "mod_def" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 865 | val dvd_def = thm "dvd_def" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 866 | val quorem_def = thm "quorem_def" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 867 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 868 | val wf_less_trans = thm "wf_less_trans"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 869 | val mod_eq = thm "mod_eq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 870 | val div_eq = thm "div_eq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 871 | val DIVISION_BY_ZERO_DIV = thm "DIVISION_BY_ZERO_DIV"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 872 | val DIVISION_BY_ZERO_MOD = thm "DIVISION_BY_ZERO_MOD"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 873 | val mod_less = thm "mod_less"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 874 | val mod_geq = thm "mod_geq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 875 | val le_mod_geq = thm "le_mod_geq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 876 | val mod_if = thm "mod_if"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 877 | val mod_1 = thm "mod_1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 878 | val mod_self = thm "mod_self"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 879 | val mod_add_self2 = thm "mod_add_self2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 880 | val mod_add_self1 = thm "mod_add_self1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 881 | val mod_mult_self1 = thm "mod_mult_self1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 882 | val mod_mult_self2 = thm "mod_mult_self2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 883 | val mod_mult_distrib = thm "mod_mult_distrib"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 884 | val mod_mult_distrib2 = thm "mod_mult_distrib2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 885 | val mod_mult_self_is_0 = thm "mod_mult_self_is_0"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 886 | val mod_mult_self1_is_0 = thm "mod_mult_self1_is_0"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 887 | val div_less = thm "div_less"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 888 | val div_geq = thm "div_geq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 889 | val le_div_geq = thm "le_div_geq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 890 | val div_if = thm "div_if"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 891 | val mod_div_equality = thm "mod_div_equality"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 892 | val mod_div_equality2 = thm "mod_div_equality2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 893 | val div_mod_equality = thm "div_mod_equality"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 894 | val div_mod_equality2 = thm "div_mod_equality2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 895 | val mult_div_cancel = thm "mult_div_cancel"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 896 | val mod_less_divisor = thm "mod_less_divisor"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 897 | val div_mult_self_is_m = thm "div_mult_self_is_m"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 898 | val div_mult_self1_is_m = thm "div_mult_self1_is_m"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 899 | val unique_quotient_lemma = thm "unique_quotient_lemma"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 900 | val unique_quotient = thm "unique_quotient"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 901 | val unique_remainder = thm "unique_remainder"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 902 | val div_0 = thm "div_0"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 903 | val mod_0 = thm "mod_0"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 904 | val div_mult1_eq = thm "div_mult1_eq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 905 | val mod_mult1_eq = thm "mod_mult1_eq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 906 | val mod_mult1_eq' = thm "mod_mult1_eq'"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 907 | val mod_mult_distrib_mod = thm "mod_mult_distrib_mod"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 908 | val div_add1_eq = thm "div_add1_eq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 909 | val mod_add1_eq = thm "mod_add1_eq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 910 | val mod_lemma = thm "mod_lemma"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 911 | val div_mult2_eq = thm "div_mult2_eq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 912 | val mod_mult2_eq = thm "mod_mult2_eq"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 913 | val div_mult_mult_lemma = thm "div_mult_mult_lemma"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 914 | val div_mult_mult1 = thm "div_mult_mult1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 915 | val div_mult_mult2 = thm "div_mult_mult2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 916 | val div_1 = thm "div_1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 917 | val div_self = thm "div_self"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 918 | val div_add_self2 = thm "div_add_self2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 919 | val div_add_self1 = thm "div_add_self1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 920 | val div_mult_self1 = thm "div_mult_self1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 921 | val div_mult_self2 = thm "div_mult_self2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 922 | val div_le_mono = thm "div_le_mono"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 923 | val div_le_mono2 = thm "div_le_mono2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 924 | val div_le_dividend = thm "div_le_dividend"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 925 | val div_less_dividend = thm "div_less_dividend"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 926 | val mod_Suc = thm "mod_Suc"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 927 | val dvdI = thm "dvdI"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 928 | val dvdE = thm "dvdE"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 929 | val dvd_0_right = thm "dvd_0_right"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 930 | val dvd_0_left = thm "dvd_0_left"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 931 | val dvd_0_left_iff = thm "dvd_0_left_iff"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 932 | val dvd_1_left = thm "dvd_1_left"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 933 | val dvd_1_iff_1 = thm "dvd_1_iff_1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 934 | val dvd_refl = thm "dvd_refl"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 935 | val dvd_trans = thm "dvd_trans"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 936 | val dvd_anti_sym = thm "dvd_anti_sym"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 937 | val dvd_add = thm "dvd_add"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 938 | val dvd_diff = thm "dvd_diff"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 939 | val dvd_diffD = thm "dvd_diffD"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 940 | val dvd_diffD1 = thm "dvd_diffD1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 941 | val dvd_mult = thm "dvd_mult"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 942 | val dvd_mult2 = thm "dvd_mult2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 943 | val dvd_reduce = thm "dvd_reduce"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 944 | val dvd_mod = thm "dvd_mod"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 945 | val dvd_mod_imp_dvd = thm "dvd_mod_imp_dvd"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 946 | val dvd_mod_iff = thm "dvd_mod_iff"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 947 | val dvd_mult_cancel = thm "dvd_mult_cancel"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 948 | val dvd_mult_cancel1 = thm "dvd_mult_cancel1"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 949 | val dvd_mult_cancel2 = thm "dvd_mult_cancel2"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 950 | val mult_dvd_mono = thm "mult_dvd_mono"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 951 | val dvd_mult_left = thm "dvd_mult_left"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 952 | val dvd_mult_right = thm "dvd_mult_right"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 953 | val dvd_imp_le = thm "dvd_imp_le"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 954 | val dvd_eq_mod_eq_0 = thm "dvd_eq_mod_eq_0"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 955 | val dvd_mult_div_cancel = thm "dvd_mult_div_cancel"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 956 | val mod_eq_0_iff = thm "mod_eq_0_iff"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 957 | val mod_eqD = thm "mod_eqD"; | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 958 | *} | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 959 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 960 | |
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 961 | (* | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 962 | lemma split_div: | 
| 13152 | 963 | assumes m: "m \<noteq> 0" | 
| 964 | shows "P(n div m :: nat) = (!i. !j<m. n = m*i + j \<longrightarrow> P i)" | |
| 965 | (is "?P = ?Q") | |
| 966 | proof | |
| 967 | assume P: ?P | |
| 968 | show ?Q | |
| 969 | proof (intro allI impI) | |
| 970 | fix i j | |
| 971 | assume n: "n = m*i + j" and j: "j < m" | |
| 972 | show "P i" | |
| 973 | proof (cases) | |
| 974 | assume "i = 0" | |
| 975 | with n j P show "P i" by simp | |
| 976 | next | |
| 977 | assume "i \<noteq> 0" | |
| 978 | with n j P show "P i" by (simp add:add_ac div_mult_self1) | |
| 979 | qed | |
| 980 | qed | |
| 981 | next | |
| 982 | assume Q: ?Q | |
| 983 | from m Q[THEN spec,of "n div m",THEN spec, of "n mod m"] | |
| 13517 | 984 | show ?P by simp | 
| 13152 | 985 | qed | 
| 986 | ||
| 987 | lemma split_mod: | |
| 988 | assumes m: "m \<noteq> 0" | |
| 989 | shows "P(n mod m :: nat) = (!i. !j<m. n = m*i + j \<longrightarrow> P j)" | |
| 990 | (is "?P = ?Q") | |
| 991 | proof | |
| 992 | assume P: ?P | |
| 993 | show ?Q | |
| 994 | proof (intro allI impI) | |
| 995 | fix i j | |
| 996 | assume "n = m*i + j" "j < m" | |
| 997 | thus "P j" using m P by(simp add:add_ac mult_ac) | |
| 998 | qed | |
| 999 | next | |
| 1000 | assume Q: ?Q | |
| 1001 | from m Q[THEN spec,of "n div m",THEN spec, of "n mod m"] | |
| 13517 | 1002 | show ?P by simp | 
| 13152 | 1003 | qed | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1004 | *) | 
| 3366 | 1005 | end |