| author | wenzelm | 
| Fri, 21 May 2004 21:14:52 +0200 | |
| changeset 14766 | c0401da7726d | 
| parent 14740 | c8e1937110c2 | 
| child 15003 | 6145dd7538d7 | 
| 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 | ||
| 14259 | 10 | theory IntDef = Equiv + NatArith: | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 11 | |
| 5508 | 12 | constdefs | 
| 14271 | 13 | intrel :: "((nat * nat) * (nat * nat)) set" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 14 |     --{*the equivalence relation underlying the integers*}
 | 
| 14496 | 15 |     "intrel == {((x,y),(u,v)) | x y u v. x+v = u+y}"
 | 
| 5508 | 16 | |
| 17 | typedef (Integ) | |
| 14259 | 18 | int = "UNIV//intrel" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 19 | by (auto simp add: quotient_def) | 
| 5508 | 20 | |
| 14691 | 21 | instance int :: "{ord, zero, one, plus, times, minus}" ..
 | 
| 5508 | 22 | |
| 23 | constdefs | |
| 14259 | 24 | int :: "nat => int" | 
| 10834 | 25 |   "int m == Abs_Integ(intrel `` {(m,0)})"
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 26 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 27 | |
| 14269 | 28 | defs (overloaded) | 
| 14271 | 29 | |
| 14259 | 30 | Zero_int_def: "0 == int 0" | 
| 14271 | 31 | One_int_def: "1 == int 1" | 
| 8937 | 32 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 33 | minus_int_def: | 
| 14532 | 34 |     "- 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 | 35 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 36 | add_int_def: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 37 | "z + w == | 
| 14532 | 38 | Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. \<Union>(u,v) \<in> Rep_Integ w. | 
| 39 | 		 intrel``{(x+u, y+v)})"
 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 40 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 41 | diff_int_def: "z - (w::int) == z + (-w)" | 
| 5508 | 42 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 43 | mult_int_def: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 44 | "z * w == | 
| 14532 | 45 | Abs_Integ (\<Union>(x,y) \<in> Rep_Integ z. \<Union>(u,v) \<in> Rep_Integ w. | 
| 46 | 		  intrel``{(x*u + y*v, x*v + y*u)})"
 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 47 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 48 | le_int_def: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 49 | "z \<le> (w::int) == | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 50 | \<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Integ z & (u,v) \<in> Rep_Integ w" | 
| 5508 | 51 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 52 | less_int_def: "(z < (w::int)) == (z \<le> w & z \<noteq> w)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 53 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 54 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 55 | subsection{*Construction of the Integers*}
 | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 56 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 57 | subsubsection{*Preliminary Lemmas about the Equivalence Relation*}
 | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 58 | |
| 14496 | 59 | 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 | 60 | by (simp add: intrel_def) | 
| 14259 | 61 | |
| 62 | lemma equiv_intrel: "equiv UNIV intrel" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 63 | by (simp add: intrel_def equiv_def refl_def sym_def trans_def) | 
| 14259 | 64 | |
| 14496 | 65 | text{*Reduces equality of equivalence classes to the @{term intrel} relation:
 | 
| 66 |   @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \<in> intrel)"} *}
 | |
| 67 | lemmas equiv_intrel_iff = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I] | |
| 14259 | 68 | |
| 14496 | 69 | declare equiv_intrel_iff [simp] | 
| 70 | ||
| 71 | ||
| 72 | text{*All equivalence classes belong to set of representatives*}
 | |
| 14532 | 73 | lemma [simp]: "intrel``{(x,y)} \<in> Integ"
 | 
| 14496 | 74 | by (auto simp add: Integ_def intrel_def quotient_def) | 
| 14259 | 75 | |
| 76 | lemma inj_on_Abs_Integ: "inj_on Abs_Integ Integ" | |
| 77 | apply (rule inj_on_inverseI) | |
| 78 | apply (erule Abs_Integ_inverse) | |
| 79 | done | |
| 80 | ||
| 14496 | 81 | text{*This theorem reduces equality on abstractions to equality on 
 | 
| 82 | representatives: | |
| 83 |   @{term "\<lbrakk>x \<in> Integ; y \<in> Integ\<rbrakk> \<Longrightarrow> (Abs_Integ x = Abs_Integ y) = (x=y)"} *}
 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 84 | declare inj_on_Abs_Integ [THEN inj_on_iff, simp] | 
| 14496 | 85 | |
| 86 | declare Abs_Integ_inverse [simp] | |
| 14259 | 87 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 88 | text{*Case analysis on the representation of an integer as an equivalence
 | 
| 14485 | 89 | class of pairs of naturals.*} | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 90 | lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 91 |      "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 92 | apply (rule Rep_Integ [of z, unfolded Integ_def, THEN quotientE]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 93 | apply (drule arg_cong [where f=Abs_Integ]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 94 | apply (auto simp add: Rep_Integ_inverse) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 95 | done | 
| 14259 | 96 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 97 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 98 | subsubsection{*@{term int}: Embedding the Naturals into the Integers*}
 | 
| 14259 | 99 | |
| 100 | lemma inj_int: "inj int" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 101 | by (simp add: inj_on_def int_def) | 
| 14259 | 102 | |
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 103 | 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 | 104 | by (fast elim!: inj_int [THEN injD]) | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 105 | |
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 106 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 107 | subsubsection{*Integer Unary Negation*}
 | 
| 14259 | 108 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 109 | lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 110 | proof - | 
| 14532 | 111 |   have "congruent intrel (\<lambda>(x,y). intrel``{(y,x)})"
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 112 | by (simp add: congruent_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 113 | thus ?thesis | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 114 | by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 115 | qed | 
| 14259 | 116 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 117 | lemma zminus_zminus: "- (- z) = (z::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 118 | by (cases z, simp add: minus) | 
| 14259 | 119 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 120 | lemma zminus_0: "- 0 = (0::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 121 | by (simp add: int_def Zero_int_def minus) | 
| 14259 | 122 | |
| 123 | ||
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 124 | subsection{*Integer Addition*}
 | 
| 14259 | 125 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 126 | lemma add: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 127 |      "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) =
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 128 |       Abs_Integ (intrel``{(x+u, y+v)})"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 129 | proof - | 
| 14658 | 130 | have "congruent2 intrel intrel | 
| 14532 | 131 |         (\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). intrel `` {(x+u, y+v)}) w) z)"
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 132 | by (simp add: congruent2_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 133 | thus ?thesis | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 134 | by (simp add: add_int_def UN_UN_split_split_eq | 
| 14658 | 135 | UN_equiv_class2 [OF equiv_intrel equiv_intrel]) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 136 | qed | 
| 14259 | 137 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 138 | lemma zminus_zadd_distrib: "- (z + w) = (- z) + (- w::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 139 | by (cases z, cases w, simp add: minus add) | 
| 14259 | 140 | |
| 141 | lemma zadd_commute: "(z::int) + w = w + z" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 142 | by (cases z, cases w, simp add: add_ac add) | 
| 14259 | 143 | |
| 144 | lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 145 | by (cases z1, cases z2, cases z3, simp add: add add_assoc) | 
| 14259 | 146 | |
| 147 | (*For AC rewriting*) | |
| 14271 | 148 | lemma zadd_left_commute: "x + (y + z) = y + ((x + z) ::int)" | 
| 14259 | 149 | apply (rule mk_left_commute [of "op +"]) | 
| 150 | apply (rule zadd_assoc) | |
| 151 | apply (rule zadd_commute) | |
| 152 | done | |
| 153 | ||
| 154 | lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute | |
| 155 | ||
| 14738 | 156 | lemmas zmult_ac = OrderedGroup.mult_ac | 
| 14271 | 157 | |
| 14259 | 158 | lemma zadd_int: "(int m) + (int n) = int (m + n)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 159 | by (simp add: int_def add) | 
| 14259 | 160 | |
| 161 | lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z" | |
| 162 | by (simp add: zadd_int zadd_assoc [symmetric]) | |
| 163 | ||
| 164 | lemma int_Suc: "int (Suc m) = 1 + (int m)" | |
| 165 | by (simp add: One_int_def zadd_int) | |
| 166 | ||
| 14738 | 167 | (*also for the instance declaration int :: comm_monoid_add*) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 168 | lemma zadd_0: "(0::int) + z = z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 169 | apply (simp add: Zero_int_def int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 170 | apply (cases z, simp add: add) | 
| 14259 | 171 | done | 
| 172 | ||
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 173 | lemma zadd_0_right: "z + (0::int) = z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 174 | by (rule trans [OF zadd_commute zadd_0]) | 
| 14259 | 175 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 176 | lemma zadd_zminus_inverse2: "(- z) + z = (0::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 177 | by (cases z, simp add: int_def Zero_int_def minus add) | 
| 14259 | 178 | |
| 179 | ||
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 180 | subsection{*Integer Multiplication*}
 | 
| 14259 | 181 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 182 | text{*Congruence property for multiplication*}
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 183 | lemma mult_congruent2: | 
| 14658 | 184 | "congruent2 intrel intrel | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 185 | (%p1 p2. (%(x,y). (%(u,v). | 
| 14532 | 186 |                     intrel``{(x*u + y*v, x*v + y*u)}) p2) p1)"
 | 
| 14259 | 187 | apply (rule equiv_intrel [THEN congruent2_commuteI]) | 
| 14532 | 188 | apply (force simp add: mult_ac, clarify) | 
| 189 | apply (simp add: congruent_def mult_ac) | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 190 | apply (rename_tac u v w x y z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 191 | 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 | 192 | apply (simp add: mult_ac, arith) | 
| 14259 | 193 | apply (simp add: add_mult_distrib [symmetric]) | 
| 194 | done | |
| 195 | ||
| 14532 | 196 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 197 | lemma mult: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 198 |      "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) =
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 199 |       Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 200 | by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2 | 
| 14658 | 201 | UN_equiv_class2 [OF equiv_intrel equiv_intrel]) | 
| 14259 | 202 | |
| 203 | lemma zmult_zminus: "(- z) * w = - (z * (w::int))" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 204 | by (cases z, cases w, simp add: minus mult add_ac) | 
| 14259 | 205 | |
| 206 | lemma zmult_commute: "(z::int) * w = w * z" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 207 | by (cases z, cases w, simp add: mult add_ac mult_ac) | 
| 14259 | 208 | |
| 209 | lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 210 | by (cases z1, cases z2, cases z3, simp add: mult add_mult_distrib2 mult_ac) | 
| 14259 | 211 | |
| 212 | 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 | 213 | by (cases z1, cases z2, cases w, simp add: add mult add_mult_distrib2 mult_ac) | 
| 14259 | 214 | |
| 215 | lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)" | |
| 216 | by (simp add: zmult_commute [of w] zadd_zmult_distrib) | |
| 217 | ||
| 218 | lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)" | |
| 14496 | 219 | by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus) | 
| 14259 | 220 | |
| 221 | lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)" | |
| 222 | by (simp add: zmult_commute [of w] zdiff_zmult_distrib) | |
| 223 | ||
| 224 | lemmas int_distrib = | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 225 | zadd_zmult_distrib zadd_zmult_distrib2 | 
| 14259 | 226 | zdiff_zmult_distrib zdiff_zmult_distrib2 | 
| 227 | ||
| 228 | lemma zmult_int: "(int m) * (int n) = int (m * n)" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 229 | by (simp add: int_def mult) | 
| 14259 | 230 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 231 | lemma zmult_1: "(1::int) * z = z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 232 | by (cases z, simp add: One_int_def int_def mult) | 
| 14259 | 233 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 234 | lemma zmult_1_right: "z * (1::int) = z" | 
| 14259 | 235 | by (rule trans [OF zmult_commute zmult_1]) | 
| 236 | ||
| 237 | ||
| 14740 | 238 | text{*The integers form a @{text comm_ring_1}*}
 | 
| 14738 | 239 | instance int :: comm_ring_1 | 
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 240 | proof | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 241 | fix i j k :: int | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 242 | 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 | 243 | show "i + j = j + i" by (simp add: zadd_commute) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 244 | show "0 + i = i" by (rule zadd_0) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 245 | show "- i + i = 0" by (rule zadd_zminus_inverse2) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 246 | 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 | 247 | 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 | 248 | show "i * j = j * i" by (rule zmult_commute) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 249 | 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 | 250 | 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 | 251 | show "0 \<noteq> (1::int)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14341diff
changeset | 252 | 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 | 253 | qed | 
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 254 | |
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 255 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 256 | subsection{*The @{text "\<le>"} Ordering*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 257 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 258 | lemma le: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 259 |   "(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 | 260 | by (force simp add: le_int_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 261 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 262 | lemma zle_refl: "w \<le> (w::int)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 263 | by (cases w, simp add: le) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 264 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 265 | 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 | 266 | 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 | 267 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 268 | 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 | 269 | 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 | 270 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 271 | (* Axiom 'order_less_le' of class 'order': *) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 272 | 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 | 273 | by (simp add: less_int_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 274 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 275 | instance int :: order | 
| 14691 | 276 | by intro_classes | 
| 277 | (assumption | | |
| 278 | 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 | 279 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 280 | (* Axiom 'linorder_linear' of class 'linorder': *) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 281 | lemma zle_linear: "(z::int) \<le> w | w \<le> z" | 
| 14691 | 282 | 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 | 283 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 284 | instance int :: linorder | 
| 14691 | 285 | by intro_classes (rule zle_linear) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 286 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 287 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 288 | lemmas zless_linear = linorder_less_linear [where 'a = int] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 289 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 290 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 291 | 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 | 292 | by (simp add: Zero_int_def) | 
| 14259 | 293 | |
| 294 | lemma zless_int [simp]: "(int m < int n) = (m<n)" | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 295 | by (simp add: le add int_def linorder_not_le [symmetric]) | 
| 14259 | 296 | |
| 297 | lemma int_less_0_conv [simp]: "~ (int k < 0)" | |
| 298 | by (simp add: Zero_int_def) | |
| 299 | ||
| 300 | lemma zero_less_int_conv [simp]: "(0 < int n) = (0 < n)" | |
| 301 | by (simp add: Zero_int_def) | |
| 302 | ||
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 303 | 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 | 304 | by (simp only: Zero_int_def One_int_def One_nat_def zless_int) | 
| 14259 | 305 | |
| 14341 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 306 | 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 | 307 | 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 | 308 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 309 | 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 | 310 | by (simp add: linorder_not_less [symmetric]) | 
| 14259 | 311 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 312 | 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 | 313 | by (simp add: Zero_int_def) | 
| 14259 | 314 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 315 | 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 | 316 | by (simp add: Zero_int_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 317 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 318 | lemma int_0 [simp]: "int 0 = (0::int)" | 
| 14259 | 319 | by (simp add: Zero_int_def) | 
| 320 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 321 | lemma int_1 [simp]: "int 1 = 1" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 322 | by (simp add: One_int_def) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 323 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 324 | 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 | 325 | 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 | 326 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 327 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 328 | subsection{*Monotonicity results*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 329 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 330 | 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 | 331 | 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 | 332 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 333 | lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 334 | apply (cases i, cases j, cases k) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 335 | apply (simp add: linorder_not_le [where 'a = int, symmetric] | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 336 | 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 | 337 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 338 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 339 | 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 | 340 | 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 | 341 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 342 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 343 | subsection{*Strict Monotonicity of Multiplication*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 344 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 345 | text{*strict, in 1st argument; proof is by induction on k>0*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 346 | lemma zmult_zless_mono2_lemma [rule_format]: | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 347 | "i<j ==> 0<k --> int k * i < int k * j" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 348 | apply (induct_tac "k", simp) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 349 | apply (simp add: int_Suc) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 350 | apply (case_tac "n=0") | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 351 | 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 | 352 | 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 | 353 | done | 
| 14259 | 354 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 355 | 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 | 356 | apply (cases k) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 357 | apply (auto simp add: le add int_def Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 358 | 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 | 359 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 360 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 361 | 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 | 362 | apply (frule order_less_imp_le [THEN zero_le_imp_eq_int]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 363 | 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 | 364 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 365 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 366 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 367 | defs (overloaded) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 368 | 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 | 369 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 370 | |
| 14740 | 371 | text{*The integers form an ordered @{text comm_ring_1}*}
 | 
| 14738 | 372 | instance int :: ordered_idom | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 373 | proof | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 374 | fix i j k :: int | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 375 | 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 | 376 | 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 | 377 | 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 | 378 | qed | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 379 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 380 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 381 | lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 382 | apply (cases w, cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 383 | 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 | 384 | done | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 385 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 386 | 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 | 387 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 388 | constdefs | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 389 | nat :: "int => nat" | 
| 14532 | 390 |     "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 | 391 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 392 | lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 393 | proof - | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 394 |   have "congruent intrel (\<lambda>(x,y). {x-y})"
 | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 395 | by (simp add: congruent_def, arith) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 396 | thus ?thesis | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 397 | by (simp add: nat_def UN_equiv_class [OF equiv_intrel]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 398 | qed | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 399 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 400 | lemma nat_int [simp]: "nat(int n) = n" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 401 | by (simp add: nat int_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 402 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 403 | lemma nat_zero [simp]: "nat 0 = 0" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 404 | 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 | 405 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 406 | 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 | 407 | by (cases z, simp add: nat le int_def Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 408 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 409 | corollary nat_0_le: "0 \<le> z ==> int (nat z) = z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 410 | apply simp | 
| 14259 | 411 | done | 
| 412 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 413 | lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 414 | by (cases z, simp add: nat le int_def Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 415 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 416 | 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 | 417 | apply (cases w, cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 418 | 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 | 419 | done | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 420 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 421 | text{*An alternative condition is @{term "0 \<le> w"} *}
 | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 422 | corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 423 | by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 424 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 425 | 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 | 426 | by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) | 
| 
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 zless_nat_conj: "(nat w < nat z) = (0 < z & w < z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 429 | apply (cases w, cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 430 | 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 | 431 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 432 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 433 | 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 | 434 | by (blast dest: nat_0_le sym) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 435 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 436 | 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 | 437 | 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 | 438 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 439 | 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 | 440 | by (simp only: eq_commute [of m] nat_eq_iff) | 
| 
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 nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 443 | apply (cases w) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 444 | 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 | 445 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 446 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 447 | 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 | 448 | by (auto simp add: nat_eq_iff2) | 
| 
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 zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 451 | by (insert zless_nat_conj [of 0], auto) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 452 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 453 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 454 | lemma nat_add_distrib: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 455 | "[| (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 | 456 | 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 | 457 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 458 | lemma nat_diff_distrib: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 459 | "[| (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 | 460 | by (cases z, cases z', | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 461 | 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 | 462 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 463 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 464 | lemma nat_zminus_int [simp]: "nat (- (int n)) = 0" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 465 | by (simp add: int_def minus nat Zero_int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 466 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 467 | lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 468 | 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 | 469 | |
| 14378 
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 | 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 | 472 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 473 | lemma negative_zless_0: "- (int (Suc n)) < 0" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 474 | by (simp add: order_less_le) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 475 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 476 | 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 | 477 | 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 | 478 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 479 | lemma negative_zle_0: "- int n \<le> 0" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 480 | by (simp add: minus_le_iff) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 481 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 482 | 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 | 483 | 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 | 484 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 485 | 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 | 486 | by (subst le_minus_iff, simp) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 487 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 488 | 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 | 489 | by (simp add: int_def le minus Zero_int_def) | 
| 14259 | 490 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 491 | 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 | 492 | by (simp add: linorder_not_less) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 493 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 494 | 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 | 495 | 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 | 496 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 497 | lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 498 | apply (cases w, cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 499 | apply (auto simp add: le add int_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 500 | apply (rename_tac a b c d) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 501 | apply (rule_tac x="c+b - (a+d)" in exI) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 502 | apply arith | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 503 | done | 
| 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" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 506 | by (simp add: zabs_def) | 
| 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 | |
| 14740 | 563 | subsection{*Embedding of the Naturals into any @{text
 | 
| 564 | comm_semiring_1_cancel}: @{term of_nat}*}
 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 565 | |
| 14738 | 566 | consts of_nat :: "nat => 'a::comm_semiring_1_cancel" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 567 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 568 | primrec | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 569 | of_nat_0: "of_nat 0 = 0" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 570 | of_nat_Suc: "of_nat (Suc m) = of_nat m + 1" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 571 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 572 | lemma of_nat_1 [simp]: "of_nat 1 = 1" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 573 | by simp | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 574 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 575 | lemma of_nat_add [simp]: "of_nat (m+n) = of_nat m + of_nat n" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 576 | apply (induct m) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 577 | apply (simp_all add: add_ac) | 
| 14378 
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 of_nat_mult [simp]: "of_nat (m*n) = of_nat m * of_nat n" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 581 | apply (induct m) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 582 | apply (simp_all add: mult_ac add_ac right_distrib) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 583 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 584 | |
| 14738 | 585 | lemma zero_le_imp_of_nat: "0 \<le> (of_nat m::'a::ordered_semidom)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 586 | apply (induct m, simp_all) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 587 | apply (erule order_trans) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 588 | apply (rule less_add_one [THEN order_less_imp_le]) | 
| 14259 | 589 | done | 
| 590 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 591 | lemma less_imp_of_nat_less: | 
| 14738 | 592 | "m < n ==> of_nat m < (of_nat n::'a::ordered_semidom)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 593 | apply (induct m n rule: diff_induct, simp_all) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 594 | apply (insert add_le_less_mono [OF zero_le_imp_of_nat zero_less_one], force) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 595 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 596 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 597 | lemma of_nat_less_imp_less: | 
| 14738 | 598 | "of_nat m < (of_nat n::'a::ordered_semidom) ==> m < n" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 599 | apply (induct m n rule: diff_induct, simp_all) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 600 | apply (insert zero_le_imp_of_nat) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 601 | apply (force simp add: linorder_not_less [symmetric]) | 
| 14259 | 602 | done | 
| 603 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 604 | lemma of_nat_less_iff [simp]: | 
| 14738 | 605 | "(of_nat m < (of_nat n::'a::ordered_semidom)) = (m<n)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 606 | by (blast intro: of_nat_less_imp_less less_imp_of_nat_less) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 607 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 608 | text{*Special cases where either operand is zero*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 609 | declare of_nat_less_iff [of 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 610 | declare of_nat_less_iff [of _ 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 611 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 612 | lemma of_nat_le_iff [simp]: | 
| 14738 | 613 | "(of_nat m \<le> (of_nat n::'a::ordered_semidom)) = (m \<le> n)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 614 | by (simp add: linorder_not_less [symmetric]) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 615 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 616 | text{*Special cases where either operand is zero*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 617 | declare of_nat_le_iff [of 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 618 | declare of_nat_le_iff [of _ 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 619 | |
| 14740 | 620 | text{*The ordering on the @{text comm_semiring_1_cancel} is necessary
 | 
| 621 | to exclude the possibility of a finite field, which indeed wraps back to | |
| 622 | zero.*} | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 623 | lemma of_nat_eq_iff [simp]: | 
| 14738 | 624 | "(of_nat m = (of_nat n::'a::ordered_semidom)) = (m = n)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 625 | by (simp add: order_eq_iff) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 626 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 627 | text{*Special cases where either operand is zero*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 628 | declare of_nat_eq_iff [of 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 629 | declare of_nat_eq_iff [of _ 0, simplified, simp] | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 630 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 631 | lemma of_nat_diff [simp]: | 
| 14738 | 632 | "n \<le> m ==> of_nat (m - n) = of_nat m - (of_nat n :: 'a::comm_ring_1)" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 633 | by (simp del: of_nat_add | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 634 | add: compare_rls of_nat_add [symmetric] split add: nat_diff_split) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 635 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 636 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 637 | subsection{*The Set of Natural Numbers*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 638 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 639 | constdefs | 
| 14738 | 640 | Nats :: "'a::comm_semiring_1_cancel set" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 641 | "Nats == range of_nat" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 642 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 643 | syntax (xsymbols)    Nats :: "'a set"   ("\<nat>")
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 644 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 645 | lemma of_nat_in_Nats [simp]: "of_nat n \<in> Nats" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 646 | by (simp add: Nats_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 647 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 648 | lemma Nats_0 [simp]: "0 \<in> Nats" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 649 | apply (simp add: Nats_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 650 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 651 | apply (rule of_nat_0 [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 652 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 653 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 654 | lemma Nats_1 [simp]: "1 \<in> Nats" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 655 | apply (simp add: Nats_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 656 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 657 | apply (rule of_nat_1 [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 658 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 659 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 660 | 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 | 661 | apply (auto simp add: Nats_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 662 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 663 | apply (rule of_nat_add [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 664 | done | 
| 
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 | 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 | 667 | apply (auto simp add: Nats_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 668 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 669 | apply (rule of_nat_mult [symmetric]) | 
| 14259 | 670 | done | 
| 671 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 672 | 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 | 673 | 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 | 674 | proof | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 675 | fix n | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 676 | 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 | 677 | qed | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 678 | |
| 14496 | 679 | lemma of_nat_eq_id [simp]: "of_nat = (id :: nat => nat)" | 
| 680 | proof | |
| 681 | fix n | |
| 682 | show "of_nat n = id n" by (induct n, simp_all) | |
| 683 | qed | |
| 684 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 685 | |
| 14740 | 686 | subsection{*Embedding of the Integers into any @{text comm_ring_1}:
 | 
| 687 | @{term of_int}*}
 | |
| 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 | constdefs | 
| 14738 | 690 | of_int :: "int => 'a::comm_ring_1" | 
| 14532 | 691 |    "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 | 692 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 693 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 694 | lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
 | 
| 14496 | 695 | proof - | 
| 696 |   have "congruent intrel (\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) })"
 | |
| 697 | by (simp add: congruent_def compare_rls of_nat_add [symmetric] | |
| 698 | del: of_nat_add) | |
| 699 | thus ?thesis | |
| 700 | by (simp add: of_int_def UN_equiv_class [OF equiv_intrel]) | |
| 701 | qed | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 702 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 703 | 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 | 704 | 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 | 705 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 706 | 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 | 707 | 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 | 708 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 709 | 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 | 710 | 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 | 711 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 712 | lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 713 | by (cases z, simp add: compare_rls of_int minus) | 
| 14259 | 714 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 715 | 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 | 716 | by (simp add: diff_minus) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 717 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 718 | 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 | 719 | apply (cases w, cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 720 | 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 | 721 | mult add_ac) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 722 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 723 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 724 | lemma of_int_le_iff [simp]: | 
| 14738 | 725 | "(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 | 726 | apply (cases w) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 727 | apply (cases z) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 728 | 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 | 729 | 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 | 730 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 731 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 732 | text{*Special cases where either operand is zero*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 733 | 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 | 734 | declare of_int_le_iff [of _ 0, simplified, simp] | 
| 14259 | 735 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 736 | lemma of_int_less_iff [simp]: | 
| 14738 | 737 | "(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 | 738 | by (simp add: linorder_not_le [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 739 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 740 | text{*Special cases where either operand is zero*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 741 | 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 | 742 | 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 | 743 | |
| 14740 | 744 | text{*The ordering on the @{text comm_ring_1} is necessary.
 | 
| 745 |  See @{text of_nat_eq_iff} above.*}
 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 746 | lemma of_int_eq_iff [simp]: | 
| 14738 | 747 | "(of_int w = (of_int z::'a::ordered_idom)) = (w = z)" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 748 | by (simp add: order_eq_iff) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 749 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 750 | text{*Special cases where either operand is zero*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 751 | 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 | 752 | 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 | 753 | |
| 14496 | 754 | lemma of_int_eq_id [simp]: "of_int = (id :: int => int)" | 
| 755 | proof | |
| 756 | fix z | |
| 757 | show "of_int z = id z" | |
| 758 | by (cases z, | |
| 759 | simp add: of_int add minus int_eq_of_nat [symmetric] int_def diff_minus) | |
| 760 | qed | |
| 761 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 762 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 763 | subsection{*The Set of Integers*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 764 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 765 | constdefs | 
| 14738 | 766 | Ints :: "'a::comm_ring_1 set" | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 767 | "Ints == range of_int" | 
| 14271 | 768 | |
| 14259 | 769 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 770 | syntax (xsymbols) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 771 |   Ints      :: "'a set"                   ("\<int>")
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 772 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 773 | lemma Ints_0 [simp]: "0 \<in> Ints" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 774 | apply (simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 775 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 776 | apply (rule of_int_0 [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 777 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 778 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 779 | lemma Ints_1 [simp]: "1 \<in> Ints" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 780 | apply (simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 781 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 782 | apply (rule of_int_1 [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 783 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 784 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 785 | 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 | 786 | apply (auto simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 787 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 788 | apply (rule of_int_add [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 789 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 790 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 791 | lemma Ints_minus [simp]: "a \<in> Ints ==> -a \<in> Ints" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 792 | apply (auto simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 793 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 794 | apply (rule of_int_minus [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 795 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 796 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 797 | 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 | 798 | apply (auto simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 799 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 800 | apply (rule of_int_diff [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 801 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 802 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 803 | 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 | 804 | apply (auto simp add: Ints_def) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 805 | apply (rule range_eqI) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 806 | apply (rule of_int_mult [symmetric]) | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 807 | done | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 808 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 809 | text{*Collapse nested embeddings*}
 | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 810 | 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 | 811 | by (induct n, auto) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 812 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 813 | lemma of_int_int_eq [simp]: "of_int (int n) = int n" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 814 | 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 | 815 | |
| 
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
 paulson parents: 
14271diff
changeset | 816 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 817 | 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 | 818 | "q \<in> \<int> ==> (!!z. q = of_int z ==> C) ==> C" | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 819 | proof (simp add: Ints_def) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 820 | assume "!!z. q = of_int z ==> C" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 821 | assume "q \<in> range of_int" thus C .. | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 822 | qed | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 823 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 824 | 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 | 825 | "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 | 826 | by (rule Ints_cases) auto | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 827 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 828 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 829 | (* int (Suc n) = 1 + int n *) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 830 | declare int_Suc [simp] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 831 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 832 | text{*Simplification of @{term "x-y < 0"}, etc.*}
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 833 | declare less_iff_diff_less_0 [symmetric, simp] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 834 | declare eq_iff_diff_eq_0 [symmetric, simp] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 835 | declare le_iff_diff_le_0 [symmetric, simp] | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 836 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 837 | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 838 | 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 | 839 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 840 | text{*By Jeremy Avigad*}
 | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 841 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 842 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 843 | lemma setsum_of_nat: "of_nat (setsum f A) = setsum (of_nat \<circ> f) A" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 844 | apply (case_tac "finite A") | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 845 | apply (erule finite_induct, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 846 | apply (simp add: setsum_def) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 847 | done | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 848 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 849 | lemma setsum_of_int: "of_int (setsum f A) = setsum (of_int \<circ> f) A" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 850 | apply (case_tac "finite A") | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 851 | apply (erule finite_induct, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 852 | apply (simp add: setsum_def) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 853 | done | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 854 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 855 | lemma int_setsum: "int (setsum f A) = setsum (int \<circ> f) A" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 856 | by (subst int_eq_of_nat, rule setsum_of_nat) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 857 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 858 | lemma setprod_of_nat: "of_nat (setprod f A) = setprod (of_nat \<circ> f) A" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 859 | apply (case_tac "finite A") | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 860 | apply (erule finite_induct, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 861 | apply (simp add: setprod_def) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 862 | done | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 863 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 864 | lemma setprod_of_int: "of_int (setprod f A) = setprod (of_int \<circ> f) A" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 865 | apply (case_tac "finite A") | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 866 | apply (erule finite_induct, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 867 | apply (simp add: setprod_def) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 868 | done | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 869 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 870 | lemma int_setprod: "int (setprod f A) = setprod (int \<circ> f) A" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 871 | by (subst int_eq_of_nat, rule setprod_of_nat) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 872 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 873 | lemma setsum_constant: "finite A ==> (\<Sum>x \<in> A. y) = of_nat(card A) * y" | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 874 | apply (erule finite_induct) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 875 | apply (auto simp add: ring_distrib add_ac) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 876 | done | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 877 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 878 | lemma setprod_nonzero_nat: | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 879 | "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 | 880 | by (rule setprod_nonzero, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 881 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 882 | lemma setprod_zero_eq_nat: | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 883 | "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 | 884 | by (rule setprod_zero_eq, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 885 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 886 | lemma setprod_nonzero_int: | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 887 | "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 | 888 | by (rule setprod_nonzero, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 889 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 890 | lemma setprod_zero_eq_int: | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 891 | "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 | 892 | by (rule setprod_zero_eq, auto) | 
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 893 | |
| 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14421diff
changeset | 894 | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 895 | 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 | 896 | whether an integer is negative or not.*} | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 897 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 898 | lemma zless_iff_Suc_zadd: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 899 | "(w < z) = (\<exists>n. z = w + int(Suc n))" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 900 | apply (cases z, cases w) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 901 | apply (auto simp add: le add int_def linorder_not_le [symmetric]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 902 | apply (rename_tac a b c d) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 903 | apply (rule_tac x="a+d - Suc(c+b)" in exI) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 904 | apply arith | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 905 | done | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 906 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 907 | lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 908 | apply (cases x) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 909 | apply (auto simp add: le minus Zero_int_def int_def order_less_le) | 
| 14496 | 910 | apply (rule_tac x="y - Suc x" in exI, arith) | 
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 911 | done | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 912 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 913 | theorem int_cases [cases type: int, case_names nonneg neg]: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 914 | "[|!! n. z = int n ==> P; !! n. z = - (int (Suc n)) ==> P |] ==> P" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 915 | apply (case_tac "z < 0", blast dest!: negD) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 916 | apply (simp add: linorder_not_less) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 917 | apply (blast dest: nat_0_le [THEN sym]) | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 918 | done | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 919 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 920 | theorem int_induct [induct type: int, case_names nonneg neg]: | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 921 | "[|!! n. P (int n); !!n. P (- (int (Suc n))) |] ==> P z" | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 922 | by (cases z) auto | 
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 923 | |
| 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 924 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 925 | (*Legacy ML bindings, but no longer the structure Int.*) | 
| 14259 | 926 | ML | 
| 927 | {*
 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 928 | val zabs_def = thm "zabs_def" | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 929 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 930 | val int_0 = thm "int_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 931 | val int_1 = thm "int_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 932 | 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 | 933 | 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 | 934 | 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 | 935 | val not_neg_0 = thm "not_neg_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 936 | val not_neg_1 = thm "not_neg_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 937 | val iszero_0 = thm "iszero_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 938 | val not_iszero_1 = thm "not_iszero_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 939 | 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 | 940 | 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 | 941 | val negative_zless = thm "negative_zless"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 942 | val negative_zle = thm "negative_zle"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 943 | 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 | 944 | 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 | 945 | val negative_eq_positive = thm "negative_eq_positive"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 946 | val zle_iff_zadd = thm "zle_iff_zadd"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 947 | val abs_int_eq = thm "abs_int_eq"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 948 | val abs_split = thm"abs_split"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 949 | val nat_int = thm "nat_int"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 950 | val nat_zminus_int = thm "nat_zminus_int"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 951 | val nat_zero = thm "nat_zero"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 952 | val not_neg_nat = thm "not_neg_nat"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 953 | val neg_nat = thm "neg_nat"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 954 | 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 | 955 | val nat_0_le = thm "nat_0_le"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 956 | val nat_le_0 = thm "nat_le_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 957 | val zless_nat_conj = thm "zless_nat_conj"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 958 | val int_cases = thm "int_cases"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 959 | |
| 14259 | 960 | val int_def = thm "int_def"; | 
| 961 | val Zero_int_def = thm "Zero_int_def"; | |
| 962 | val One_int_def = thm "One_int_def"; | |
| 14479 
0eca4aabf371
streamlined treatment of quotients for the integers
 paulson parents: 
14430diff
changeset | 963 | val diff_int_def = thm "diff_int_def"; | 
| 14259 | 964 | |
| 965 | val inj_int = thm "inj_int"; | |
| 966 | val zminus_zminus = thm "zminus_zminus"; | |
| 967 | val zminus_0 = thm "zminus_0"; | |
| 968 | val zminus_zadd_distrib = thm "zminus_zadd_distrib"; | |
| 969 | val zadd_commute = thm "zadd_commute"; | |
| 970 | val zadd_assoc = thm "zadd_assoc"; | |
| 971 | val zadd_left_commute = thm "zadd_left_commute"; | |
| 972 | val zadd_ac = thms "zadd_ac"; | |
| 14271 | 973 | val zmult_ac = thms "zmult_ac"; | 
| 14259 | 974 | val zadd_int = thm "zadd_int"; | 
| 975 | val zadd_int_left = thm "zadd_int_left"; | |
| 976 | val int_Suc = thm "int_Suc"; | |
| 977 | val zadd_0 = thm "zadd_0"; | |
| 978 | val zadd_0_right = thm "zadd_0_right"; | |
| 979 | val zmult_zminus = thm "zmult_zminus"; | |
| 980 | val zmult_commute = thm "zmult_commute"; | |
| 981 | val zmult_assoc = thm "zmult_assoc"; | |
| 982 | val zadd_zmult_distrib = thm "zadd_zmult_distrib"; | |
| 983 | val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2"; | |
| 984 | val zdiff_zmult_distrib = thm "zdiff_zmult_distrib"; | |
| 985 | val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2"; | |
| 986 | val int_distrib = thms "int_distrib"; | |
| 987 | val zmult_int = thm "zmult_int"; | |
| 988 | val zmult_1 = thm "zmult_1"; | |
| 989 | val zmult_1_right = thm "zmult_1_right"; | |
| 990 | val int_int_eq = thm "int_int_eq"; | |
| 991 | val int_eq_0_conv = thm "int_eq_0_conv"; | |
| 992 | val zless_int = thm "zless_int"; | |
| 993 | val int_less_0_conv = thm "int_less_0_conv"; | |
| 994 | val zero_less_int_conv = thm "zero_less_int_conv"; | |
| 995 | val zle_int = thm "zle_int"; | |
| 996 | val zero_zle_int = thm "zero_zle_int"; | |
| 997 | val int_le_0_conv = thm "int_le_0_conv"; | |
| 998 | val zle_refl = thm "zle_refl"; | |
| 999 | val zle_linear = thm "zle_linear"; | |
| 1000 | val zle_trans = thm "zle_trans"; | |
| 1001 | 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 | 1002 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1003 | val Ints_def = thm "Ints_def"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1004 | val Nats_def = thm "Nats_def"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1005 | |
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1006 | val of_nat_0 = thm "of_nat_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1007 | val of_nat_Suc = thm "of_nat_Suc"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1008 | val of_nat_1 = thm "of_nat_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1009 | val of_nat_add = thm "of_nat_add"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1010 | val of_nat_mult = thm "of_nat_mult"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1011 | 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 | 1012 | 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 | 1013 | 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 | 1014 | 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 | 1015 | 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 | 1016 | 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 | 1017 | val Nats_0 = thm "Nats_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1018 | val Nats_1 = thm "Nats_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1019 | val Nats_add = thm "Nats_add"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1020 | val Nats_mult = thm "Nats_mult"; | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 1021 | 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 | 1022 | val of_int = thm "of_int"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1023 | val of_int_0 = thm "of_int_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1024 | val of_int_1 = thm "of_int_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1025 | val of_int_add = thm "of_int_add"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1026 | val of_int_minus = thm "of_int_minus"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1027 | val of_int_diff = thm "of_int_diff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1028 | val of_int_mult = thm "of_int_mult"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1029 | 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 | 1030 | 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 | 1031 | 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 | 1032 | val Ints_0 = thm "Ints_0"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1033 | val Ints_1 = thm "Ints_1"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1034 | val Ints_add = thm "Ints_add"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1035 | val Ints_minus = thm "Ints_minus"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1036 | val Ints_diff = thm "Ints_diff"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1037 | val Ints_mult = thm "Ints_mult"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1038 | 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 | 1039 | val Ints_cases = thm "Ints_cases"; | 
| 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14348diff
changeset | 1040 | val Ints_induct = thm "Ints_induct"; | 
| 14259 | 1041 | *} | 
| 1042 | ||
| 5508 | 1043 | end |