| author | haftmann | 
| Sun, 05 Jun 2005 14:41:23 +0200 | |
| changeset 16276 | 3a50bf1f04d0 | 
| parent 15921 | b6e345548913 | 
| child 16413 | 47ffc49c7d7b | 
| permissions | -rw-r--r-- | 
| 5508 | 1 | (* Title: IntDef.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1996 University of Cambridge | |
| 5 | ||
| 6 | *) | |
| 7 | ||
| 14271 | 8 | header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*}
 | 
| 9 | ||
| 15131 | 10 | theory IntDef | 
| 15300 | 11 | imports Equiv_Relations NatArith | 
| 15131 | 12 | begin | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 13 | |
| 5508 | 14 | constdefs | 
| 14271 | 15 | intrel :: "((nat * nat) * (nat * nat)) set" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 16 |     --{*the equivalence relation underlying the integers*}
 | 
| 14496 | 17 |     "intrel == {((x,y),(u,v)) | x y u v. x+v = u+y}"
 | 
| 5508 | 18 | |
| 19 | typedef (Integ) | |
| 14259 | 20 | int = "UNIV//intrel" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 21 | by (auto simp add: quotient_def) | 
| 5508 | 22 | |
| 14691 | 23 | instance int :: "{ord, zero, one, plus, times, minus}" ..
 | 
| 5508 | 24 | |
| 25 | constdefs | |
| 14259 | 26 | int :: "nat => int" | 
| 10834 | 27 |   "int m == Abs_Integ(intrel `` {(m,0)})"
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 28 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 29 | |
| 14269 | 30 | defs (overloaded) | 
| 14271 | 31 | |
| 14259 | 32 | Zero_int_def: "0 == int 0" | 
| 14271 | 33 | One_int_def: "1 == int 1" | 
| 8937 | 34 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 35 | minus_int_def: | 
| 14532 | 36 |     "- z == Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. intrel``{(y,x)})"
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 37 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 38 | add_int_def: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 39 | "z + w == | 
| 14532 | 40 | Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. \<Union>(u,v) \<in> Rep_Integ w. | 
| 41 | 		 intrel``{(x+u, y+v)})"
 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 42 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 43 | diff_int_def: "z - (w::int) == z + (-w)" | 
| 5508 | 44 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 45 | mult_int_def: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 46 | "z * w == | 
| 14532 | 47 | Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. \<Union>(u,v) \<in> Rep_Integ w. | 
| 48 | 		  intrel``{(x*u + y*v, x*v + y*u)})"
 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 49 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 50 | le_int_def: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 51 | "z \<le> (w::int) == | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 52 | \<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Integ z & (u,v) \<in> Rep_Integ w" | 
| 5508 | 53 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 54 | less_int_def: "(z < (w::int)) == (z \<le> w & z \<noteq> w)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 55 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 56 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 57 | subsection{*Construction of the Integers*}
 | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 58 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 59 | subsubsection{*Preliminary Lemmas about the Equivalence Relation*}
 | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 60 | |
| 14496 | 61 | lemma intrel_iff [simp]: "(((x,y),(u,v)) \<in> intrel) = (x+v = u+y)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 62 | by (simp add: intrel_def) | 
| 14259 | 63 | |
| 64 | lemma equiv_intrel: "equiv UNIV intrel" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 65 | by (simp add: intrel_def equiv_def refl_def sym_def trans_def) | 
| 14259 | 66 | |
| 14496 | 67 | text{*Reduces equality of equivalence classes to the @{term intrel} relation:
 | 
| 68 |   @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \<in> intrel)"} *}
 | |
| 69 | lemmas equiv_intrel_iff = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I] | |
| 14259 | 70 | |
| 14496 | 71 | declare equiv_intrel_iff [simp] | 
| 72 | ||
| 73 | ||
| 74 | text{*All equivalence classes belong to set of representatives*}
 | |
| 14532 | 75 | lemma [simp]: "intrel``{(x,y)} \<in> Integ"
 | 
| 14496 | 76 | by (auto simp add: Integ_def intrel_def quotient_def) | 
| 14259 | 77 | |
| 15413 | 78 | text{*Reduces equality on abstractions to equality on representatives:
 | 
| 14496 | 79 |   @{term "\<lbrakk>x \<in> Integ; y \<in> Integ\<rbrakk> \<Longrightarrow> (Abs_Integ x = Abs_Integ y) = (x=y)"} *}
 | 
| 15413 | 80 | declare Abs_Integ_inject [simp] Abs_Integ_inverse [simp] | 
| 14259 | 81 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 82 | text{*Case analysis on the representation of an integer as an equivalence
 | 
| 14485 | 83 | class of pairs of naturals.*} | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 84 | lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 85 |      "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
 | 
| 15413 | 86 | apply (rule Abs_Integ_cases [of z]) | 
| 87 | apply (auto simp add: Integ_def quotient_def) | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 88 | done | 
| 14259 | 89 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 90 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 91 | subsubsection{*@{term int}: Embedding the Naturals into the Integers*}
 | 
| 14259 | 92 | |
| 93 | lemma inj_int: "inj int" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 94 | by (simp add: inj_on_def int_def) | 
| 14259 | 95 | |
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 96 | lemma int_int_eq [iff]: "(int m = int n) = (m = n)" | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 97 | by (fast elim!: inj_int [THEN injD]) | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 98 | |
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 99 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 100 | subsubsection{*Integer Unary Negation*}
 | 
| 14259 | 101 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 102 | lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 103 | proof - | 
| 15169 | 104 |   have "(\<lambda>(x,y). intrel``{(y,x)}) respects intrel"
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 105 | by (simp add: congruent_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 106 | thus ?thesis | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 107 | by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 108 | qed | 
| 14259 | 109 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 110 | lemma zminus_zminus: "- (- z) = (z::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 111 | by (cases z, simp add: minus) | 
| 14259 | 112 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 113 | lemma zminus_0: "- 0 = (0::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 114 | by (simp add: int_def Zero_int_def minus) | 
| 14259 | 115 | |
| 116 | ||
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 117 | subsection{*Integer Addition*}
 | 
| 14259 | 118 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 119 | lemma add: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 120 |      "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) =
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 121 |       Abs_Integ (intrel``{(x+u, y+v)})"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 122 | proof - | 
| 15169 | 123 |   have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). intrel `` {(x+u, y+v)}) w) z) 
 | 
| 124 | respects2 intrel" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 125 | by (simp add: congruent2_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 126 | thus ?thesis | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 127 | by (simp add: add_int_def UN_UN_split_split_eq | 
| 14658 | 128 | UN_equiv_class2 [OF equiv_intrel equiv_intrel]) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 129 | qed | 
| 14259 | 130 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 131 | lemma zminus_zadd_distrib: "- (z + w) = (- z) + (- w::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 132 | by (cases z, cases w, simp add: minus add) | 
| 14259 | 133 | |
| 134 | lemma zadd_commute: "(z::int) + w = w + z" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 135 | by (cases z, cases w, simp add: add_ac add) | 
| 14259 | 136 | |
| 137 | lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 138 | by (cases z1, cases z2, cases z3, simp add: add add_assoc) | 
| 14259 | 139 | |
| 140 | (*For AC rewriting*) | |
| 14271 | 141 | lemma zadd_left_commute: "x + (y + z) = y + ((x + z) ::int)" | 
| 14259 | 142 | apply (rule mk_left_commute [of "op +"]) | 
| 143 | apply (rule zadd_assoc) | |
| 144 | apply (rule zadd_commute) | |
| 145 | done | |
| 146 | ||
| 147 | lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute | |
| 148 | ||
| 14738 | 149 | lemmas zmult_ac = OrderedGroup.mult_ac | 
| 14271 | 150 | |
| 14259 | 151 | lemma zadd_int: "(int m) + (int n) = int (m + n)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 152 | by (simp add: int_def add) | 
| 14259 | 153 | |
| 154 | lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z" | |
| 155 | by (simp add: zadd_int zadd_assoc [symmetric]) | |
| 156 | ||
| 157 | lemma int_Suc: "int (Suc m) = 1 + (int m)" | |
| 158 | by (simp add: One_int_def zadd_int) | |
| 159 | ||
| 14738 | 160 | (*also for the instance declaration int :: comm_monoid_add*) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 161 | lemma zadd_0: "(0::int) + z = z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 162 | apply (simp add: Zero_int_def int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 163 | apply (cases z, simp add: add) | 
| 14259 | 164 | done | 
| 165 | ||
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 166 | lemma zadd_0_right: "z + (0::int) = z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 167 | by (rule trans [OF zadd_commute zadd_0]) | 
| 14259 | 168 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 169 | lemma zadd_zminus_inverse2: "(- z) + z = (0::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 170 | by (cases z, simp add: int_def Zero_int_def minus add) | 
| 14259 | 171 | |
| 172 | ||
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 173 | subsection{*Integer Multiplication*}
 | 
| 14259 | 174 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 175 | text{*Congruence property for multiplication*}
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 176 | lemma mult_congruent2: | 
| 15169 | 177 |      "(%p1 p2. (%(x,y). (%(u,v). intrel``{(x*u + y*v, x*v + y*u)}) p2) p1)
 | 
| 178 | respects2 intrel" | |
| 14259 | 179 | apply (rule equiv_intrel [THEN congruent2_commuteI]) | 
| 14532 | 180 | apply (force simp add: mult_ac, clarify) | 
| 181 | apply (simp add: congruent_def mult_ac) | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 182 | apply (rename_tac u v w x y z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 183 | apply (subgoal_tac "u*y + x*y = w*y + v*y & u*z + x*z = w*z + v*z") | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 184 | apply (simp add: mult_ac, arith) | 
| 14259 | 185 | apply (simp add: add_mult_distrib [symmetric]) | 
| 186 | done | |
| 187 | ||
| 14532 | 188 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 189 | lemma mult: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 190 |      "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) =
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 191 |       Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 192 | by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2 | 
| 14658 | 193 | UN_equiv_class2 [OF equiv_intrel equiv_intrel]) | 
| 14259 | 194 | |
| 195 | lemma zmult_zminus: "(- z) * w = - (z * (w::int))" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 196 | by (cases z, cases w, simp add: minus mult add_ac) | 
| 14259 | 197 | |
| 198 | lemma zmult_commute: "(z::int) * w = w * z" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 199 | by (cases z, cases w, simp add: mult add_ac mult_ac) | 
| 14259 | 200 | |
| 201 | lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 202 | by (cases z1, cases z2, cases z3, simp add: mult add_mult_distrib2 mult_ac) | 
| 14259 | 203 | |
| 204 | lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 205 | by (cases z1, cases z2, cases w, simp add: add mult add_mult_distrib2 mult_ac) | 
| 14259 | 206 | |
| 207 | lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)" | |
| 208 | by (simp add: zmult_commute [of w] zadd_zmult_distrib) | |
| 209 | ||
| 210 | lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)" | |
| 14496 | 211 | by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus) | 
| 14259 | 212 | |
| 213 | lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)" | |
| 214 | by (simp add: zmult_commute [of w] zdiff_zmult_distrib) | |
| 215 | ||
| 216 | lemmas int_distrib = | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 217 | zadd_zmult_distrib zadd_zmult_distrib2 | 
| 14259 | 218 | zdiff_zmult_distrib zdiff_zmult_distrib2 | 
| 219 | ||
| 220 | lemma zmult_int: "(int m) * (int n) = int (m * n)" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 221 | by (simp add: int_def mult) | 
| 14259 | 222 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 223 | lemma zmult_1: "(1::int) * z = z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 224 | by (cases z, simp add: One_int_def int_def mult) | 
| 14259 | 225 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 226 | lemma zmult_1_right: "z * (1::int) = z" | 
| 14259 | 227 | by (rule trans [OF zmult_commute zmult_1]) | 
| 228 | ||
| 229 | ||
| 14740 | 230 | text{*The integers form a @{text comm_ring_1}*}
 | 
| 14738 | 231 | instance int :: comm_ring_1 | 
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 232 | proof | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 233 | fix i j k :: int | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 234 | show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc) | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 235 | show "i + j = j + i" by (simp add: zadd_commute) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 236 | show "0 + i = i" by (rule zadd_0) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 237 | show "- i + i = 0" by (rule zadd_zminus_inverse2) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 238 | show "i - j = i + (-j)" by (simp add: diff_int_def) | 
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 239 | show "(i * j) * k = i * (j * k)" by (rule zmult_assoc) | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 240 | show "i * j = j * i" by (rule zmult_commute) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 241 | show "1 * i = i" by (rule zmult_1) | 
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 242 | show "(i + j) * k = i * k + j * k" by (simp add: int_distrib) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 243 | show "0 \<noteq> (1::int)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 244 | by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq) | 
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 245 | qed | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 246 | |
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 247 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 248 | subsection{*The @{text "\<le>"} Ordering*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 249 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 250 | lemma le: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 251 |   "(Abs_Integ(intrel``{(x,y)}) \<le> Abs_Integ(intrel``{(u,v)})) = (x+v \<le> u+y)"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 252 | by (force simp add: le_int_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 253 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 254 | lemma zle_refl: "w \<le> (w::int)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 255 | by (cases w, simp add: le) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 256 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 257 | lemma zle_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::int)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 258 | by (cases i, cases j, cases k, simp add: le) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 259 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 260 | lemma zle_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::int)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 261 | by (cases w, cases z, simp add: le) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 262 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 263 | (* Axiom 'order_less_le' of class 'order': *) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 264 | lemma zless_le: "((w::int) < z) = (w \<le> z & w \<noteq> z)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 265 | by (simp add: less_int_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 266 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 267 | instance int :: order | 
| 14691 | 268 | by intro_classes | 
| 269 | (assumption | | |
| 270 | rule zle_refl zle_trans zle_anti_sym zless_le)+ | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 271 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 272 | (* Axiom 'linorder_linear' of class 'linorder': *) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 273 | lemma zle_linear: "(z::int) \<le> w | w \<le> z" | 
| 14691 | 274 | by (cases z, cases w) (simp add: le linorder_linear) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 275 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 276 | instance int :: linorder | 
| 14691 | 277 | by intro_classes (rule zle_linear) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 278 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 279 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 280 | lemmas zless_linear = linorder_less_linear [where 'a = int] | 
| 15921 | 281 | lemmas linorder_neqE_int = linorder_neqE[where 'a = int] | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 282 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 283 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 284 | lemma int_eq_0_conv [simp]: "(int n = 0) = (n = 0)" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 285 | by (simp add: Zero_int_def) | 
| 14259 | 286 | |
| 287 | lemma zless_int [simp]: "(int m < int n) = (m<n)" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 288 | by (simp add: le add int_def linorder_not_le [symmetric]) | 
| 14259 | 289 | |
| 290 | lemma int_less_0_conv [simp]: "~ (int k < 0)" | |
| 291 | by (simp add: Zero_int_def) | |
| 292 | ||
| 293 | lemma zero_less_int_conv [simp]: "(0 < int n) = (0 < n)" | |
| 294 | by (simp add: Zero_int_def) | |
| 295 | ||
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 296 | lemma int_0_less_1: "0 < (1::int)" | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 297 | by (simp only: Zero_int_def One_int_def One_nat_def zless_int) | 
| 14259 | 298 | |
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 299 | lemma int_0_neq_1 [simp]: "0 \<noteq> (1::int)" | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 300 | by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq) | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 301 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 302 | lemma zle_int [simp]: "(int m \<le> int n) = (m\<le>n)" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 303 | by (simp add: linorder_not_less [symmetric]) | 
| 14259 | 304 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 305 | lemma zero_zle_int [simp]: "(0 \<le> int n)" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 306 | by (simp add: Zero_int_def) | 
| 14259 | 307 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 308 | lemma int_le_0_conv [simp]: "(int n \<le> 0) = (n = 0)" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 309 | by (simp add: Zero_int_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 310 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 311 | lemma int_0 [simp]: "int 0 = (0::int)" | 
| 14259 | 312 | by (simp add: Zero_int_def) | 
| 313 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 314 | lemma int_1 [simp]: "int 1 = 1" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 315 | by (simp add: One_int_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 316 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 317 | lemma int_Suc0_eq_1: "int (Suc 0) = 1" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 318 | by (simp add: One_int_def One_nat_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 319 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 320 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 321 | subsection{*Monotonicity results*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 322 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 323 | lemma zadd_left_mono: "i \<le> j ==> k + i \<le> k + (j::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 324 | by (cases i, cases j, cases k, simp add: le add) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 325 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 326 | lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 327 | apply (cases i, cases j, cases k) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 328 | apply (simp add: linorder_not_le [where 'a = int, symmetric] | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 329 | linorder_not_le [where 'a = nat] le add) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 330 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 331 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 332 | lemma zadd_zless_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::int)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 333 | by (rule order_less_le_trans [OF zadd_strict_right_mono zadd_left_mono]) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 334 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 335 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 336 | subsection{*Strict Monotonicity of Multiplication*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 337 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 338 | text{*strict, in 1st argument; proof is by induction on k>0*}
 | 
| 15251 | 339 | lemma zmult_zless_mono2_lemma: | 
| 340 | "i<j ==> 0<k ==> int k * i < int k * j" | |
| 341 | apply (induct "k", simp) | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 342 | apply (simp add: int_Suc) | 
| 15251 | 343 | apply (case_tac "k=0") | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 344 | apply (simp_all add: zadd_zmult_distrib int_Suc0_eq_1 order_le_less) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 345 | apply (simp add: zadd_zless_mono int_Suc0_eq_1 order_le_less) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 346 | done | 
| 14259 | 347 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 348 | lemma zero_le_imp_eq_int: "0 \<le> k ==> \<exists>n. k = int n" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 349 | apply (cases k) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 350 | apply (auto simp add: le add int_def Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 351 | apply (rule_tac x="x-y" in exI, simp) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 352 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 353 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 354 | lemma zmult_zless_mono2: "[| i<j; (0::int) < k |] ==> k*i < k*j" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 355 | apply (frule order_less_imp_le [THEN zero_le_imp_eq_int]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 356 | apply (auto simp add: zmult_zless_mono2_lemma) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 357 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 358 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 359 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 360 | defs (overloaded) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 361 | zabs_def: "abs(i::int) == if i < 0 then -i else i" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 362 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 363 | |
| 14740 | 364 | text{*The integers form an ordered @{text comm_ring_1}*}
 | 
| 14738 | 365 | instance int :: ordered_idom | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 366 | proof | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 367 | fix i j k :: int | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 368 | show "i \<le> j ==> k + i \<le> k + j" by (rule zadd_left_mono) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 369 | show "i < j ==> 0 < k ==> k * i < k * j" by (rule zmult_zless_mono2) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 370 | show "\<bar>i\<bar> = (if i < 0 then -i else i)" by (simp only: zabs_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 371 | qed | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 372 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 373 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 374 | lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 375 | apply (cases w, cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 376 | apply (simp add: linorder_not_le [symmetric] le int_def add One_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 377 | done | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 378 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 379 | subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 380 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 381 | constdefs | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 382 | nat :: "int => nat" | 
| 14532 | 383 |     "nat z == contents (\<Union>(x,y) \<in> Rep_Integ z. { x-y })"
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 384 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 385 | lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 386 | proof - | 
| 15169 | 387 |   have "(\<lambda>(x,y). {x-y}) respects intrel"
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 388 | by (simp add: congruent_def, arith) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 389 | thus ?thesis | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 390 | by (simp add: nat_def UN_equiv_class [OF equiv_intrel]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 391 | qed | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 392 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 393 | lemma nat_int [simp]: "nat(int n) = n" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 394 | by (simp add: nat int_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 395 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 396 | lemma nat_zero [simp]: "nat 0 = 0" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 397 | by (simp only: Zero_int_def nat_int) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 398 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 399 | lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 400 | by (cases z, simp add: nat le int_def Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 401 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 402 | corollary nat_0_le: "0 \<le> z ==> int (nat z) = z" | 
| 15413 | 403 | by simp | 
| 14259 | 404 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 405 | lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 406 | by (cases z, simp add: nat le int_def Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 407 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 408 | lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 409 | apply (cases w, cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 410 | apply (simp add: nat le linorder_not_le [symmetric] int_def Zero_int_def, arith) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 411 | done | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 412 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 413 | text{*An alternative condition is @{term "0 \<le> w"} *}
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 414 | corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 415 | by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 416 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 417 | corollary nat_less_eq_zless: "0 \<le> w ==> (nat w < nat z) = (w<z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 418 | by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 419 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 420 | lemma zless_nat_conj: "(nat w < nat z) = (0 < z & w < z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 421 | apply (cases w, cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 422 | apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 423 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 424 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 425 | lemma nonneg_eq_int: "[| 0 \<le> z; !!m. z = int m ==> P |] ==> P" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 426 | by (blast dest: nat_0_le sym) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 427 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 428 | lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 429 | by (cases w, simp add: nat le int_def Zero_int_def, arith) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 430 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 431 | corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 432 | by (simp only: eq_commute [of m] nat_eq_iff) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 433 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 434 | lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 435 | apply (cases w) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 436 | apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 437 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 438 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 439 | lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \<le> z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 440 | by (auto simp add: nat_eq_iff2) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 441 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 442 | lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 443 | by (insert zless_nat_conj [of 0], auto) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 444 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 445 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 446 | lemma nat_add_distrib: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 447 | "[| (0::int) \<le> z; 0 \<le> z' |] ==> nat (z+z') = nat z + nat z'" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 448 | by (cases z, cases z', simp add: nat add le int_def Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 449 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 450 | lemma nat_diff_distrib: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 451 | "[| (0::int) \<le> z'; z' \<le> z |] ==> nat (z-z') = nat z - nat z'" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 452 | by (cases z, cases z', | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 453 | simp add: nat add minus diff_minus le int_def Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 454 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 455 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 456 | lemma nat_zminus_int [simp]: "nat (- (int n)) = 0" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 457 | by (simp add: int_def minus nat Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 458 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 459 | lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 460 | by (cases z, simp add: nat le int_def linorder_not_le [symmetric], arith) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 461 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 462 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 463 | subsection{*Lemmas about the Function @{term int} and Orderings*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 464 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 465 | lemma negative_zless_0: "- (int (Suc n)) < 0" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 466 | by (simp add: order_less_le) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 467 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 468 | lemma negative_zless [iff]: "- (int (Suc n)) < int m" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 469 | by (rule negative_zless_0 [THEN order_less_le_trans], simp) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 470 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 471 | lemma negative_zle_0: "- int n \<le> 0" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 472 | by (simp add: minus_le_iff) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 473 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 474 | lemma negative_zle [iff]: "- int n \<le> int m" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 475 | by (rule order_trans [OF negative_zle_0 zero_zle_int]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 476 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 477 | lemma not_zle_0_negative [simp]: "~ (0 \<le> - (int (Suc n)))" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 478 | by (subst le_minus_iff, simp) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 479 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 480 | lemma int_zle_neg: "(int n \<le> - int m) = (n = 0 & m = 0)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 481 | by (simp add: int_def le minus Zero_int_def) | 
| 14259 | 482 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 483 | lemma not_int_zless_negative [simp]: "~ (int n < - int m)" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 484 | by (simp add: linorder_not_less) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 485 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 486 | lemma negative_eq_positive [simp]: "(- int n = int m) = (n = 0 & m = 0)" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 487 | by (force simp add: order_eq_iff [of "- int n"] int_zle_neg) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 488 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 489 | lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)" | 
| 15413 | 490 | proof (cases w, cases z, simp add: le add int_def) | 
| 491 | fix a b c d | |
| 492 |   assume "w = Abs_Integ (intrel `` {(a,b)})" "z = Abs_Integ (intrel `` {(c,d)})"
 | |
| 493 | show "(a+d \<le> c+b) = (\<exists>n. c+b = a+n+d)" | |
| 494 | proof | |
| 495 | assume "a + d \<le> c + b" | |
| 496 | thus "\<exists>n. c + b = a + n + d" | |
| 497 | by (auto intro!: exI [where x="c+b - (a+d)"]) | |
| 498 | next | |
| 499 | assume "\<exists>n. c + b = a + n + d" | |
| 500 | then obtain n where "c + b = a + n + d" .. | |
| 501 | thus "a + d \<le> c + b" by arith | |
| 502 | qed | |
| 503 | qed | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 504 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 505 | lemma abs_int_eq [simp]: "abs (int m) = int m" | 
| 15003 | 506 | by (simp add: abs_if) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 507 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 508 | text{*This version is proved for all ordered rings, not just integers!
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 509 |       It is proved here because attribute @{text arith_split} is not available
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 510 |       in theory @{text Ring_and_Field}.
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 511 |       But is it really better than just rewriting with @{text abs_if}?*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 512 | lemma abs_split [arith_split]: | 
| 14738 | 513 | "P(abs(a::'a::ordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 514 | by (force dest: order_less_le_trans simp add: abs_if linorder_not_less) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 515 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 516 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 517 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 518 | subsection{*The Constants @{term neg} and @{term iszero}*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 519 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 520 | constdefs | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 521 | |
| 14738 | 522 | neg :: "'a::ordered_idom => bool" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 523 | "neg(Z) == Z < 0" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 524 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 525 | (*For simplifying equalities*) | 
| 14738 | 526 | iszero :: "'a::comm_semiring_1_cancel => bool" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 527 | "iszero z == z = (0)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 528 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 529 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 530 | lemma not_neg_int [simp]: "~ neg(int n)" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 531 | by (simp add: neg_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 532 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 533 | lemma neg_zminus_int [simp]: "neg(- (int (Suc n)))" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 534 | by (simp add: neg_def neg_less_0_iff_less) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 535 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 536 | lemmas neg_eq_less_0 = neg_def | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 537 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 538 | lemma not_neg_eq_ge_0: "(~neg x) = (0 \<le> x)" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 539 | by (simp add: neg_def linorder_not_less) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 540 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 541 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 542 | subsection{*To simplify inequalities when Numeral1 can get simplified to 1*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 543 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 544 | lemma not_neg_0: "~ neg 0" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 545 | by (simp add: One_int_def neg_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 546 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 547 | lemma not_neg_1: "~ neg 1" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 548 | by (simp add: neg_def linorder_not_less zero_le_one) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 549 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 550 | lemma iszero_0: "iszero 0" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 551 | by (simp add: iszero_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 552 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 553 | lemma not_iszero_1: "~ iszero 1" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 554 | by (simp add: iszero_def eq_commute) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 555 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 556 | lemma neg_nat: "neg z ==> nat z = 0" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 557 | by (simp add: neg_def order_less_imp_le) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 558 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 559 | lemma not_neg_nat: "~ neg z ==> int (nat z) = z" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 560 | by (simp add: linorder_not_less neg_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 561 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 562 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 563 | subsection{*The Set of Natural Numbers*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 564 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 565 | constdefs | 
| 14738 | 566 | Nats :: "'a::comm_semiring_1_cancel set" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 567 | "Nats == range of_nat" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 568 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 569 | syntax (xsymbols)    Nats :: "'a set"   ("\<nat>")
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 570 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 571 | lemma of_nat_in_Nats [simp]: "of_nat n \<in> Nats" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 572 | by (simp add: Nats_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 573 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 574 | lemma Nats_0 [simp]: "0 \<in> Nats" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 575 | apply (simp add: Nats_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 576 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 577 | apply (rule of_nat_0 [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 578 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 579 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 580 | lemma Nats_1 [simp]: "1 \<in> Nats" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 581 | apply (simp add: Nats_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 582 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 583 | apply (rule of_nat_1 [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 584 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 585 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 586 | lemma Nats_add [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a+b \<in> Nats" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 587 | apply (auto simp add: Nats_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 588 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 589 | apply (rule of_nat_add [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 590 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 591 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 592 | lemma Nats_mult [simp]: "[|a \<in> Nats; b \<in> Nats|] ==> a*b \<in> Nats" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 593 | apply (auto simp add: Nats_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 594 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 595 | apply (rule of_nat_mult [symmetric]) | 
| 14259 | 596 | done | 
| 597 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 598 | text{*Agreement with the specific embedding for the integers*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 599 | lemma int_eq_of_nat: "int = (of_nat :: nat => int)" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 600 | proof | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 601 | fix n | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 602 | show "int n = of_nat n" by (induct n, simp_all add: int_Suc add_ac) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 603 | qed | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 604 | |
| 14496 | 605 | lemma of_nat_eq_id [simp]: "of_nat = (id :: nat => nat)" | 
| 606 | proof | |
| 607 | fix n | |
| 608 | show "of_nat n = id n" by (induct n, simp_all) | |
| 609 | qed | |
| 610 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 611 | |
| 14740 | 612 | subsection{*Embedding of the Integers into any @{text comm_ring_1}:
 | 
| 613 | @{term of_int}*}
 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 614 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 615 | constdefs | 
| 14738 | 616 | of_int :: "int => 'a::comm_ring_1" | 
| 14532 | 617 |    "of_int z == contents (\<Union>(i,j) \<in> Rep_Integ z. { of_nat i - of_nat j })"
 | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 618 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 619 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 620 | lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
 | 
| 14496 | 621 | proof - | 
| 15169 | 622 |   have "(\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) }) respects intrel"
 | 
| 14496 | 623 | by (simp add: congruent_def compare_rls of_nat_add [symmetric] | 
| 624 | del: of_nat_add) | |
| 625 | thus ?thesis | |
| 626 | by (simp add: of_int_def UN_equiv_class [OF equiv_intrel]) | |
| 627 | qed | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 628 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 629 | lemma of_int_0 [simp]: "of_int 0 = 0" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 630 | by (simp add: of_int Zero_int_def int_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 631 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 632 | lemma of_int_1 [simp]: "of_int 1 = 1" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 633 | by (simp add: of_int One_int_def int_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 634 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 635 | lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 636 | by (cases w, cases z, simp add: compare_rls of_int add) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 637 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 638 | lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 639 | by (cases z, simp add: compare_rls of_int minus) | 
| 14259 | 640 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 641 | lemma of_int_diff [simp]: "of_int (w-z) = of_int w - of_int z" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 642 | by (simp add: diff_minus) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 643 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 644 | lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 645 | apply (cases w, cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 646 | apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 647 | mult add_ac) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 648 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 649 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 650 | lemma of_int_le_iff [simp]: | 
| 14738 | 651 | "(of_int w \<le> (of_int z::'a::ordered_idom)) = (w \<le> z)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 652 | apply (cases w) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 653 | apply (cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 654 | apply (simp add: compare_rls of_int le diff_int_def add minus | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 655 | of_nat_add [symmetric] del: of_nat_add) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 656 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 657 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 658 | text{*Special cases where either operand is zero*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 659 | declare of_int_le_iff [of 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 660 | declare of_int_le_iff [of _ 0, simplified, simp] | 
| 14259 | 661 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 662 | lemma of_int_less_iff [simp]: | 
| 14738 | 663 | "(of_int w < (of_int z::'a::ordered_idom)) = (w < z)" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 664 | by (simp add: linorder_not_le [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 665 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 666 | text{*Special cases where either operand is zero*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 667 | declare of_int_less_iff [of 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 668 | declare of_int_less_iff [of _ 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 669 | |
| 14740 | 670 | text{*The ordering on the @{text comm_ring_1} is necessary.
 | 
| 671 |  See @{text of_nat_eq_iff} above.*}
 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 672 | lemma of_int_eq_iff [simp]: | 
| 14738 | 673 | "(of_int w = (of_int z::'a::ordered_idom)) = (w = z)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 674 | by (simp add: order_eq_iff) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 675 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 676 | text{*Special cases where either operand is zero*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 677 | declare of_int_eq_iff [of 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 678 | declare of_int_eq_iff [of _ 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 679 | |
| 14496 | 680 | lemma of_int_eq_id [simp]: "of_int = (id :: int => int)" | 
| 681 | proof | |
| 682 | fix z | |
| 683 | show "of_int z = id z" | |
| 684 | by (cases z, | |
| 685 | simp add: of_int add minus int_eq_of_nat [symmetric] int_def diff_minus) | |
| 686 | qed | |
| 687 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 688 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 689 | subsection{*The Set of Integers*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 690 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 691 | constdefs | 
| 14738 | 692 | Ints :: "'a::comm_ring_1 set" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 693 | "Ints == range of_int" | 
| 14271 | 694 | |
| 14259 | 695 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 696 | syntax (xsymbols) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 697 |   Ints      :: "'a set"                   ("\<int>")
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 698 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 699 | lemma Ints_0 [simp]: "0 \<in> Ints" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 700 | apply (simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 701 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 702 | apply (rule of_int_0 [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 703 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 704 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 705 | lemma Ints_1 [simp]: "1 \<in> Ints" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 706 | apply (simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 707 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 708 | apply (rule of_int_1 [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 709 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 710 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 711 | lemma Ints_add [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a+b \<in> Ints" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 712 | apply (auto simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 713 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 714 | apply (rule of_int_add [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 715 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 716 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 717 | lemma Ints_minus [simp]: "a \<in> Ints ==> -a \<in> Ints" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 718 | apply (auto simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 719 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 720 | apply (rule of_int_minus [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 721 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 722 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 723 | lemma Ints_diff [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a-b \<in> Ints" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 724 | apply (auto simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 725 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 726 | apply (rule of_int_diff [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 727 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 728 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 729 | lemma Ints_mult [simp]: "[|a \<in> Ints; b \<in> Ints|] ==> a*b \<in> Ints" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 730 | apply (auto simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 731 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 732 | apply (rule of_int_mult [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 733 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 734 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 735 | text{*Collapse nested embeddings*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 736 | lemma of_int_of_nat_eq [simp]: "of_int (of_nat n) = of_nat n" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 737 | by (induct n, auto) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 738 | |
| 15013 | 739 | lemma of_int_int_eq [simp]: "of_int (int n) = of_nat n" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 740 | by (simp add: int_eq_of_nat) | 
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 741 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 742 | lemma Ints_cases [case_names of_int, cases set: Ints]: | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 743 | "q \<in> \<int> ==> (!!z. q = of_int z ==> C) ==> C" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 744 | proof (simp add: Ints_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 745 | assume "!!z. q = of_int z ==> C" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 746 | assume "q \<in> range of_int" thus C .. | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 747 | qed | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 748 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 749 | lemma Ints_induct [case_names of_int, induct set: Ints]: | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 750 | "q \<in> \<int> ==> (!!z. P (of_int z)) ==> P q" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 751 | by (rule Ints_cases) auto | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 752 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 753 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 754 | (* int (Suc n) = 1 + int n *) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 755 | declare int_Suc [simp] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 756 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 757 | text{*Simplification of @{term "x-y < 0"}, etc.*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 758 | declare less_iff_diff_less_0 [symmetric, simp] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 759 | declare eq_iff_diff_eq_0 [symmetric, simp] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 760 | declare le_iff_diff_le_0 [symmetric, simp] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 761 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 762 | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 763 | subsection{*More Properties of @{term setsum} and  @{term setprod}*}
 | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 764 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 765 | text{*By Jeremy Avigad*}
 | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 766 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 767 | |
| 15554 | 768 | lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 769 | apply (case_tac "finite A") | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 770 | apply (erule finite_induct, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 771 | done | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 772 | |
| 15554 | 773 | lemma of_int_setsum: "of_int (setsum f A) = (\<Sum>x\<in>A. of_int(f x))" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 774 | apply (case_tac "finite A") | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 775 | apply (erule finite_induct, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 776 | done | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 777 | |
| 15554 | 778 | lemma int_setsum: "int (setsum f A) = (\<Sum>x\<in>A. int(f x))" | 
| 779 | by (simp add: int_eq_of_nat of_nat_setsum) | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 780 | |
| 15554 | 781 | lemma of_nat_setprod: "of_nat (setprod f A) = (\<Prod>x\<in>A. of_nat(f x))" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 782 | apply (case_tac "finite A") | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 783 | apply (erule finite_induct, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 784 | done | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 785 | |
| 15554 | 786 | lemma of_int_setprod: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 787 | apply (case_tac "finite A") | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 788 | apply (erule finite_induct, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 789 | done | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 790 | |
| 15554 | 791 | lemma int_setprod: "int (setprod f A) = (\<Prod>x\<in>A. int(f x))" | 
| 792 | by (simp add: int_eq_of_nat of_nat_setprod) | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 793 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 794 | lemma setprod_nonzero_nat: | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 795 | "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::nat)) ==> setprod f A \<noteq> 0" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 796 | by (rule setprod_nonzero, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 797 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 798 | lemma setprod_zero_eq_nat: | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 799 | "finite A ==> (setprod f A = (0::nat)) = (\<exists>x \<in> A. f x = 0)" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 800 | by (rule setprod_zero_eq, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 801 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 802 | lemma setprod_nonzero_int: | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 803 | "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::int)) ==> setprod f A \<noteq> 0" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 804 | by (rule setprod_nonzero, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 805 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 806 | lemma setprod_zero_eq_int: | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 807 | "finite A ==> (setprod f A = (0::int)) = (\<exists>x \<in> A. f x = 0)" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 808 | by (rule setprod_zero_eq, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 809 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 810 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 811 | text{*Now we replace the case analysis rule by a more conventional one:
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 812 | whether an integer is negative or not.*} | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 813 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 814 | lemma zless_iff_Suc_zadd: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 815 | "(w < z) = (\<exists>n. z = w + int(Suc n))" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 816 | apply (cases z, cases w) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 817 | apply (auto simp add: le add int_def linorder_not_le [symmetric]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 818 | apply (rename_tac a b c d) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 819 | apply (rule_tac x="a+d - Suc(c+b)" in exI) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 820 | apply arith | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 821 | done | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 822 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 823 | lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 824 | apply (cases x) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 825 | apply (auto simp add: le minus Zero_int_def int_def order_less_le) | 
| 14496 | 826 | apply (rule_tac x="y - Suc x" in exI, arith) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 827 | done | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 828 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 829 | theorem int_cases [cases type: int, case_names nonneg neg]: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 830 | "[|!! n. z = int n ==> P; !! n. z = - (int (Suc n)) ==> P |] ==> P" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 831 | apply (case_tac "z < 0", blast dest!: negD) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 832 | apply (simp add: linorder_not_less) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 833 | apply (blast dest: nat_0_le [THEN sym]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 834 | done | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 835 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 836 | theorem int_induct [induct type: int, case_names nonneg neg]: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 837 | "[|!! n. P (int n); !!n. P (- (int (Suc n))) |] ==> P z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 838 | by (cases z) auto | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 839 | |
| 15558 | 840 | text{*Contributed by Brian Huffman*}
 | 
| 841 | theorem int_diff_cases [case_names diff]: | |
| 842 | assumes prem: "!!m n. z = int m - int n ==> P" shows "P" | |
| 843 | apply (rule_tac z=z in int_cases) | |
| 844 | apply (rule_tac m=n and n=0 in prem, simp) | |
| 845 | apply (rule_tac m=0 and n="Suc n" in prem, simp) | |
| 846 | done | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 847 | |
| 15013 | 848 | lemma of_nat_nat: "0 \<le> z ==> of_nat (nat z) = of_int z" | 
| 849 | apply (cases z) | |
| 850 | apply (simp_all add: not_zle_0_negative del: int_Suc) | |
| 851 | done | |
| 852 | ||
| 853 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 854 | (*Legacy ML bindings, but no longer the structure Int.*) | 
| 14259 | 855 | ML | 
| 856 | {*
 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 857 | val zabs_def = thm "zabs_def" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 858 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 859 | val int_0 = thm "int_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 860 | val int_1 = thm "int_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 861 | val int_Suc0_eq_1 = thm "int_Suc0_eq_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 862 | val neg_eq_less_0 = thm "neg_eq_less_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 863 | val not_neg_eq_ge_0 = thm "not_neg_eq_ge_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 864 | val not_neg_0 = thm "not_neg_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 865 | val not_neg_1 = thm "not_neg_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 866 | val iszero_0 = thm "iszero_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 867 | val not_iszero_1 = thm "not_iszero_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 868 | val int_0_less_1 = thm "int_0_less_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 869 | val int_0_neq_1 = thm "int_0_neq_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 870 | val negative_zless = thm "negative_zless"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 871 | val negative_zle = thm "negative_zle"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 872 | val not_zle_0_negative = thm "not_zle_0_negative"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 873 | val not_int_zless_negative = thm "not_int_zless_negative"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 874 | val negative_eq_positive = thm "negative_eq_positive"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 875 | val zle_iff_zadd = thm "zle_iff_zadd"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 876 | val abs_int_eq = thm "abs_int_eq"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 877 | val abs_split = thm"abs_split"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 878 | val nat_int = thm "nat_int"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 879 | val nat_zminus_int = thm "nat_zminus_int"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 880 | val nat_zero = thm "nat_zero"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 881 | val not_neg_nat = thm "not_neg_nat"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 882 | val neg_nat = thm "neg_nat"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 883 | val zless_nat_eq_int_zless = thm "zless_nat_eq_int_zless"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 884 | val nat_0_le = thm "nat_0_le"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 885 | val nat_le_0 = thm "nat_le_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 886 | val zless_nat_conj = thm "zless_nat_conj"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 887 | val int_cases = thm "int_cases"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 888 | |
| 14259 | 889 | val int_def = thm "int_def"; | 
| 890 | val Zero_int_def = thm "Zero_int_def"; | |
| 891 | val One_int_def = thm "One_int_def"; | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 892 | val diff_int_def = thm "diff_int_def"; | 
| 14259 | 893 | |
| 894 | val inj_int = thm "inj_int"; | |
| 895 | val zminus_zminus = thm "zminus_zminus"; | |
| 896 | val zminus_0 = thm "zminus_0"; | |
| 897 | val zminus_zadd_distrib = thm "zminus_zadd_distrib"; | |
| 898 | val zadd_commute = thm "zadd_commute"; | |
| 899 | val zadd_assoc = thm "zadd_assoc"; | |
| 900 | val zadd_left_commute = thm "zadd_left_commute"; | |
| 901 | val zadd_ac = thms "zadd_ac"; | |
| 14271 | 902 | val zmult_ac = thms "zmult_ac"; | 
| 14259 | 903 | val zadd_int = thm "zadd_int"; | 
| 904 | val zadd_int_left = thm "zadd_int_left"; | |
| 905 | val int_Suc = thm "int_Suc"; | |
| 906 | val zadd_0 = thm "zadd_0"; | |
| 907 | val zadd_0_right = thm "zadd_0_right"; | |
| 908 | val zmult_zminus = thm "zmult_zminus"; | |
| 909 | val zmult_commute = thm "zmult_commute"; | |
| 910 | val zmult_assoc = thm "zmult_assoc"; | |
| 911 | val zadd_zmult_distrib = thm "zadd_zmult_distrib"; | |
| 912 | val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2"; | |
| 913 | val zdiff_zmult_distrib = thm "zdiff_zmult_distrib"; | |
| 914 | val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2"; | |
| 915 | val int_distrib = thms "int_distrib"; | |
| 916 | val zmult_int = thm "zmult_int"; | |
| 917 | val zmult_1 = thm "zmult_1"; | |
| 918 | val zmult_1_right = thm "zmult_1_right"; | |
| 919 | val int_int_eq = thm "int_int_eq"; | |
| 920 | val int_eq_0_conv = thm "int_eq_0_conv"; | |
| 921 | val zless_int = thm "zless_int"; | |
| 922 | val int_less_0_conv = thm "int_less_0_conv"; | |
| 923 | val zero_less_int_conv = thm "zero_less_int_conv"; | |
| 924 | val zle_int = thm "zle_int"; | |
| 925 | val zero_zle_int = thm "zero_zle_int"; | |
| 926 | val int_le_0_conv = thm "int_le_0_conv"; | |
| 927 | val zle_refl = thm "zle_refl"; | |
| 928 | val zle_linear = thm "zle_linear"; | |
| 929 | val zle_trans = thm "zle_trans"; | |
| 930 | val zle_anti_sym = thm "zle_anti_sym"; | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 931 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 932 | val Ints_def = thm "Ints_def"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 933 | val Nats_def = thm "Nats_def"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 934 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 935 | val of_nat_0 = thm "of_nat_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 936 | val of_nat_Suc = thm "of_nat_Suc"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 937 | val of_nat_1 = thm "of_nat_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 938 | val of_nat_add = thm "of_nat_add"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 939 | val of_nat_mult = thm "of_nat_mult"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 940 | val zero_le_imp_of_nat = thm "zero_le_imp_of_nat"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 941 | val less_imp_of_nat_less = thm "less_imp_of_nat_less"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 942 | val of_nat_less_imp_less = thm "of_nat_less_imp_less"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 943 | val of_nat_less_iff = thm "of_nat_less_iff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 944 | val of_nat_le_iff = thm "of_nat_le_iff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 945 | val of_nat_eq_iff = thm "of_nat_eq_iff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 946 | val Nats_0 = thm "Nats_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 947 | val Nats_1 = thm "Nats_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 948 | val Nats_add = thm "Nats_add"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 949 | val Nats_mult = thm "Nats_mult"; | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 950 | val int_eq_of_nat = thm"int_eq_of_nat"; | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 951 | val of_int = thm "of_int"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 952 | val of_int_0 = thm "of_int_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 953 | val of_int_1 = thm "of_int_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 954 | val of_int_add = thm "of_int_add"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 955 | val of_int_minus = thm "of_int_minus"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 956 | val of_int_diff = thm "of_int_diff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 957 | val of_int_mult = thm "of_int_mult"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 958 | val of_int_le_iff = thm "of_int_le_iff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 959 | val of_int_less_iff = thm "of_int_less_iff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 960 | val of_int_eq_iff = thm "of_int_eq_iff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 961 | val Ints_0 = thm "Ints_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 962 | val Ints_1 = thm "Ints_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 963 | val Ints_add = thm "Ints_add"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 964 | val Ints_minus = thm "Ints_minus"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 965 | val Ints_diff = thm "Ints_diff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 966 | val Ints_mult = thm "Ints_mult"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 967 | val of_int_of_nat_eq = thm"of_int_of_nat_eq"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 968 | val Ints_cases = thm "Ints_cases"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 969 | val Ints_induct = thm "Ints_induct"; | 
| 14259 | 970 | *} | 
| 971 | ||
| 5508 | 972 | end |