| author | wenzelm | 
| Wed, 07 Nov 2007 22:20:11 +0100 | |
| changeset 25332 | 73491e84ead1 | 
| parent 25230 | 022029099a83 | 
| child 25349 | 0d46bea01741 | 
| permissions | -rw-r--r-- | 
| 23164 | 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 | ||
| 8 | header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*} 
 | |
| 9 | ||
| 10 | theory IntDef | |
| 11 | imports Equiv_Relations Nat | |
| 12 | begin | |
| 13 | ||
| 14 | text {* the equivalence relation underlying the integers *}
 | |
| 15 | ||
| 16 | definition | |
| 17 | intrel :: "((nat \<times> nat) \<times> (nat \<times> nat)) set" | |
| 18 | where | |
| 19 |   "intrel = {((x, y), (u, v)) | x y u v. x + v = u +y }"
 | |
| 20 | ||
| 21 | typedef (Integ) | |
| 22 | int = "UNIV//intrel" | |
| 23 | by (auto simp add: quotient_def) | |
| 24 | ||
| 25 | instance int :: zero | |
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 26 |   Zero_int_def: "0 \<equiv> Abs_Integ (intrel `` {(0, 0)})" ..
 | 
| 23164 | 27 | |
| 28 | instance int :: one | |
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 29 |   One_int_def: "1 \<equiv> Abs_Integ (intrel `` {(1, 0)})" ..
 | 
| 23164 | 30 | |
| 31 | instance int :: plus | |
| 32 | add_int_def: "z + w \<equiv> Abs_Integ | |
| 33 | (\<Union>(x, y) \<in> Rep_Integ z. \<Union>(u, v) \<in> Rep_Integ w. | |
| 34 |       intrel `` {(x + u, y + v)})" ..
 | |
| 35 | ||
| 36 | instance int :: minus | |
| 37 | minus_int_def: | |
| 38 |     "- z \<equiv> Abs_Integ (\<Union>(x, y) \<in> Rep_Integ z. intrel `` {(y, x)})"
 | |
| 39 | diff_int_def: "z - w \<equiv> z + (-w)" .. | |
| 40 | ||
| 41 | instance int :: times | |
| 42 | mult_int_def: "z * w \<equiv> Abs_Integ | |
| 43 | (\<Union>(x, y) \<in> Rep_Integ z. \<Union>(u,v ) \<in> Rep_Integ w. | |
| 44 |       intrel `` {(x*u + y*v, x*v + y*u)})" ..
 | |
| 45 | ||
| 46 | instance int :: ord | |
| 47 | le_int_def: | |
| 48 | "z \<le> w \<equiv> \<exists>x y u v. x+v \<le> u+y \<and> (x, y) \<in> Rep_Integ z \<and> (u, v) \<in> Rep_Integ w" | |
| 49 | less_int_def: "z < w \<equiv> z \<le> w \<and> z \<noteq> w" .. | |
| 50 | ||
| 51 | lemmas [code func del] = Zero_int_def One_int_def add_int_def | |
| 52 | minus_int_def mult_int_def le_int_def less_int_def | |
| 53 | ||
| 54 | ||
| 55 | subsection{*Construction of the Integers*}
 | |
| 56 | ||
| 57 | lemma intrel_iff [simp]: "(((x,y),(u,v)) \<in> intrel) = (x+v = u+y)" | |
| 58 | by (simp add: intrel_def) | |
| 59 | ||
| 60 | lemma equiv_intrel: "equiv UNIV intrel" | |
| 61 | by (simp add: intrel_def equiv_def refl_def sym_def trans_def) | |
| 62 | ||
| 63 | text{*Reduces equality of equivalence classes to the @{term intrel} relation:
 | |
| 64 |   @{term "(intrel `` {x} = intrel `` {y}) = ((x,y) \<in> intrel)"} *}
 | |
| 65 | lemmas equiv_intrel_iff [simp] = eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I] | |
| 66 | ||
| 67 | text{*All equivalence classes belong to set of representatives*}
 | |
| 68 | lemma [simp]: "intrel``{(x,y)} \<in> Integ"
 | |
| 69 | by (auto simp add: Integ_def intrel_def quotient_def) | |
| 70 | ||
| 71 | text{*Reduces equality on abstractions to equality on representatives:
 | |
| 72 |   @{prop "\<lbrakk>x \<in> Integ; y \<in> Integ\<rbrakk> \<Longrightarrow> (Abs_Integ x = Abs_Integ y) = (x=y)"} *}
 | |
| 24286 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 paulson parents: 
24196diff
changeset | 73 | declare Abs_Integ_inject [simp,noatp] Abs_Integ_inverse [simp,noatp] | 
| 23164 | 74 | |
| 75 | text{*Case analysis on the representation of an integer as an equivalence
 | |
| 76 | class of pairs of naturals.*} | |
| 77 | lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]: | |
| 78 |      "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P"
 | |
| 79 | apply (rule Abs_Integ_cases [of z]) | |
| 80 | apply (auto simp add: Integ_def quotient_def) | |
| 81 | done | |
| 82 | ||
| 83 | ||
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 84 | subsection{*Arithmetic Operations*}
 | 
| 23164 | 85 | |
| 86 | lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})"
 | |
| 87 | proof - | |
| 88 |   have "(\<lambda>(x,y). intrel``{(y,x)}) respects intrel"
 | |
| 89 | by (simp add: congruent_def) | |
| 90 | thus ?thesis | |
| 91 | by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel]) | |
| 92 | qed | |
| 93 | ||
| 94 | lemma add: | |
| 95 |      "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) =
 | |
| 96 |       Abs_Integ (intrel``{(x+u, y+v)})"
 | |
| 97 | proof - | |
| 98 |   have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). intrel `` {(x+u, y+v)}) w) z) 
 | |
| 99 | respects2 intrel" | |
| 100 | by (simp add: congruent2_def) | |
| 101 | thus ?thesis | |
| 102 | by (simp add: add_int_def UN_UN_split_split_eq | |
| 103 | UN_equiv_class2 [OF equiv_intrel equiv_intrel]) | |
| 104 | qed | |
| 105 | ||
| 106 | text{*Congruence property for multiplication*}
 | |
| 107 | lemma mult_congruent2: | |
| 108 |      "(%p1 p2. (%(x,y). (%(u,v). intrel``{(x*u + y*v, x*v + y*u)}) p2) p1)
 | |
| 109 | respects2 intrel" | |
| 110 | apply (rule equiv_intrel [THEN congruent2_commuteI]) | |
| 111 | apply (force simp add: mult_ac, clarify) | |
| 112 | apply (simp add: congruent_def mult_ac) | |
| 113 | apply (rename_tac u v w x y z) | |
| 114 | apply (subgoal_tac "u*y + x*y = w*y + v*y & u*z + x*z = w*z + v*z") | |
| 115 | apply (simp add: mult_ac) | |
| 116 | apply (simp add: add_mult_distrib [symmetric]) | |
| 117 | done | |
| 118 | ||
| 119 | lemma mult: | |
| 120 |      "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) =
 | |
| 121 |       Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})"
 | |
| 122 | by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2 | |
| 123 | UN_equiv_class2 [OF equiv_intrel equiv_intrel]) | |
| 124 | ||
| 125 | text{*The integers form a @{text comm_ring_1}*}
 | |
| 126 | instance int :: comm_ring_1 | |
| 127 | proof | |
| 128 | fix i j k :: int | |
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 129 | show "(i + j) + k = i + (j + k)" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 130 | by (cases i, cases j, cases k) (simp add: add add_assoc) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 131 | show "i + j = j + i" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 132 | by (cases i, cases j) (simp add: add_ac add) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 133 | show "0 + i = i" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 134 | by (cases i) (simp add: Zero_int_def add) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 135 | show "- i + i = 0" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 136 | by (cases i) (simp add: Zero_int_def minus add) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 137 | show "i - j = i + - j" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 138 | by (simp add: diff_int_def) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 139 | show "(i * j) * k = i * (j * k)" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23438diff
changeset | 140 | by (cases i, cases j, cases k) (simp add: mult ring_simps) | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 141 | show "i * j = j * i" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23438diff
changeset | 142 | by (cases i, cases j) (simp add: mult ring_simps) | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 143 | show "1 * i = i" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 144 | by (cases i) (simp add: One_int_def mult) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 145 | show "(i + j) * k = i * k + j * k" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23438diff
changeset | 146 | by (cases i, cases j, cases k) (simp add: add mult ring_simps) | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 147 | show "0 \<noteq> (1::int)" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 148 | by (simp add: Zero_int_def One_int_def) | 
| 23164 | 149 | qed | 
| 150 | ||
| 24196 | 151 | lemma int_def: "of_nat m = Abs_Integ (intrel `` {(m, 0)})"
 | 
| 23365 | 152 | by (induct m, simp_all add: Zero_int_def One_int_def add) | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 153 | |
| 23164 | 154 | |
| 155 | subsection{*The @{text "\<le>"} Ordering*}
 | |
| 156 | ||
| 157 | lemma le: | |
| 158 |   "(Abs_Integ(intrel``{(x,y)}) \<le> Abs_Integ(intrel``{(u,v)})) = (x+v \<le> u+y)"
 | |
| 159 | by (force simp add: le_int_def) | |
| 160 | ||
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 161 | lemma less: | 
| 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 162 |   "(Abs_Integ(intrel``{(x,y)}) < Abs_Integ(intrel``{(u,v)})) = (x+v < u+y)"
 | 
| 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 163 | by (simp add: less_int_def le order_less_le) | 
| 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 164 | |
| 23164 | 165 | instance int :: linorder | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 166 | proof | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 167 | fix i j k :: int | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 168 | show "(i < j) = (i \<le> j \<and> i \<noteq> j)" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 169 | by (simp add: less_int_def) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 170 | show "i \<le> i" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 171 | by (cases i) (simp add: le) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 172 | show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> i \<le> k" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 173 | by (cases i, cases j, cases k) (simp add: le) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 174 | show "i \<le> j \<Longrightarrow> j \<le> i \<Longrightarrow> i = j" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 175 | by (cases i, cases j) (simp add: le) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 176 | show "i \<le> j \<or> j \<le> i" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 177 | by (cases i, cases j) (simp add: le linorder_linear) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 178 | qed | 
| 23164 | 179 | |
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 180 | instance int :: pordered_cancel_ab_semigroup_add | 
| 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 181 | proof | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 182 | fix i j k :: int | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 183 | show "i \<le> j \<Longrightarrow> k + i \<le> k + j" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 184 | by (cases i, cases j, cases k) (simp add: le add) | 
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 185 | qed | 
| 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 186 | |
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 187 | text{*Strict Monotonicity of Multiplication*}
 | 
| 23164 | 188 | |
| 189 | text{*strict, in 1st argument; proof is by induction on k>0*}
 | |
| 190 | lemma zmult_zless_mono2_lemma: | |
| 24196 | 191 | "(i::int)<j ==> 0<k ==> of_nat k * i < of_nat k * j" | 
| 23164 | 192 | apply (induct "k", simp) | 
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 193 | apply (simp add: left_distrib) | 
| 23164 | 194 | apply (case_tac "k=0") | 
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 195 | apply (simp_all add: add_strict_mono) | 
| 23164 | 196 | done | 
| 197 | ||
| 24196 | 198 | lemma zero_le_imp_eq_int: "(0::int) \<le> k ==> \<exists>n. k = of_nat n" | 
| 23164 | 199 | apply (cases k) | 
| 23365 | 200 | apply (auto simp add: le add int_def Zero_int_def) | 
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 201 | apply (rule_tac x="x-y" in exI, simp) | 
| 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 202 | done | 
| 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 203 | |
| 24196 | 204 | lemma zero_less_imp_eq_int: "(0::int) < k ==> \<exists>n>0. k = of_nat n" | 
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 205 | apply (cases k) | 
| 23365 | 206 | apply (simp add: less int_def Zero_int_def) | 
| 23164 | 207 | apply (rule_tac x="x-y" in exI, simp) | 
| 208 | done | |
| 209 | ||
| 210 | lemma zmult_zless_mono2: "[| i<j; (0::int) < k |] ==> k*i < k*j" | |
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 211 | apply (drule zero_less_imp_eq_int) | 
| 23164 | 212 | apply (auto simp add: zmult_zless_mono2_lemma) | 
| 213 | done | |
| 214 | ||
| 23879 | 215 | instance int :: abs | 
| 23164 | 216 | zabs_def: "\<bar>i\<Colon>int\<bar> \<equiv> if i < 0 then - i else i" .. | 
| 24506 | 217 | instance int :: sgn | 
| 218 | zsgn_def: "sgn(i\<Colon>int) \<equiv> (if i=0 then 0 else if 0<i then 1 else - 1)" .. | |
| 23164 | 219 | |
| 220 | instance int :: distrib_lattice | |
| 221 | "inf \<equiv> min" | |
| 222 | "sup \<equiv> max" | |
| 223 | by intro_classes | |
| 224 | (auto simp add: inf_int_def sup_int_def min_max.sup_inf_distrib1) | |
| 225 | ||
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 226 | text{*The integers form an ordered integral domain*}
 | 
| 23164 | 227 | instance int :: ordered_idom | 
| 228 | proof | |
| 229 | fix i j k :: int | |
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 230 | show "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 231 | by (rule zmult_zless_mono2) | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 232 | show "\<bar>i\<bar> = (if i < 0 then -i else i)" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 233 | by (simp only: zabs_def) | 
| 24506 | 234 | show "sgn(i\<Colon>int) = (if i=0 then 0 else if 0<i then 1 else - 1)" | 
| 235 | by (simp only: zsgn_def) | |
| 23164 | 236 | qed | 
| 237 | ||
| 238 | lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z" | |
| 239 | apply (cases w, cases z) | |
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 240 | apply (simp add: less le add One_int_def) | 
| 23164 | 241 | done | 
| 242 | ||
| 23299 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 243 | |
| 
292b1cbd05dc
remove dependencies of proofs on constant int::nat=>int, preparing to remove it
 huffman parents: 
23282diff
changeset | 244 | subsection{*Magnitude of an Integer, as a Natural Number: @{term nat}*}
 | 
| 23164 | 245 | |
| 246 | definition | |
| 247 | nat :: "int \<Rightarrow> nat" | |
| 248 | where | |
| 249 |   [code func del]: "nat z = contents (\<Union>(x, y) \<in> Rep_Integ z. {x-y})"
 | |
| 250 | ||
| 251 | lemma nat: "nat (Abs_Integ (intrel``{(x,y)})) = x-y"
 | |
| 252 | proof - | |
| 253 |   have "(\<lambda>(x,y). {x-y}) respects intrel"
 | |
| 254 | by (simp add: congruent_def) arith | |
| 255 | thus ?thesis | |
| 256 | by (simp add: nat_def UN_equiv_class [OF equiv_intrel]) | |
| 257 | qed | |
| 258 | ||
| 24196 | 259 | lemma nat_int [simp]: "nat (of_nat n) = n" | 
| 23365 | 260 | by (simp add: nat int_def) | 
| 23164 | 261 | |
| 262 | lemma nat_zero [simp]: "nat 0 = 0" | |
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 263 | by (simp add: Zero_int_def nat) | 
| 23164 | 264 | |
| 24196 | 265 | lemma int_nat_eq [simp]: "of_nat (nat z) = (if 0 \<le> z then z else 0)" | 
| 23365 | 266 | by (cases z, simp add: nat le int_def Zero_int_def) | 
| 23164 | 267 | |
| 24196 | 268 | corollary nat_0_le: "0 \<le> z ==> of_nat (nat z) = z" | 
| 23164 | 269 | by simp | 
| 270 | ||
| 271 | lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0" | |
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 272 | by (cases z, simp add: nat le Zero_int_def) | 
| 23164 | 273 | |
| 274 | lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)" | |
| 275 | apply (cases w, cases z) | |
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 276 | apply (simp add: nat le linorder_not_le [symmetric] Zero_int_def, arith) | 
| 23164 | 277 | done | 
| 278 | ||
| 279 | text{*An alternative condition is @{term "0 \<le> w"} *}
 | |
| 280 | corollary nat_mono_iff: "0 < z ==> (nat w < nat z) = (w < z)" | |
| 281 | by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) | |
| 282 | ||
| 283 | corollary nat_less_eq_zless: "0 \<le> w ==> (nat w < nat z) = (w<z)" | |
| 284 | by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) | |
| 285 | ||
| 23365 | 286 | lemma zless_nat_conj [simp]: "(nat w < nat z) = (0 < z & w < z)" | 
| 23164 | 287 | apply (cases w, cases z) | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 288 | apply (simp add: nat le Zero_int_def linorder_not_le [symmetric], arith) | 
| 23164 | 289 | done | 
| 290 | ||
| 24196 | 291 | lemma nonneg_eq_int: | 
| 292 | fixes z :: int | |
| 293 | assumes "0 \<le> z" and "\<And>m. z = of_nat m \<Longrightarrow> P" | |
| 294 | shows P | |
| 295 | using assms by (blast dest: nat_0_le sym) | |
| 23164 | 296 | |
| 24196 | 297 | lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = of_nat m else m=0)" | 
| 23365 | 298 | by (cases w, simp add: nat le int_def Zero_int_def, arith) | 
| 23164 | 299 | |
| 24196 | 300 | corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = of_nat m else m=0)" | 
| 23365 | 301 | by (simp only: eq_commute [of m] nat_eq_iff) | 
| 23164 | 302 | |
| 24196 | 303 | lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < of_nat m)" | 
| 23164 | 304 | apply (cases w) | 
| 23365 | 305 | apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith) | 
| 23164 | 306 | done | 
| 307 | ||
| 24196 | 308 | lemma int_eq_iff: "(of_nat m = z) = (m = nat z & 0 \<le> z)" | 
| 23365 | 309 | by (auto simp add: nat_eq_iff2) | 
| 23164 | 310 | |
| 311 | lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)" | |
| 312 | by (insert zless_nat_conj [of 0], auto) | |
| 313 | ||
| 314 | lemma nat_add_distrib: | |
| 315 | "[| (0::int) \<le> z; 0 \<le> z' |] ==> nat (z+z') = nat z + nat z'" | |
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 316 | by (cases z, cases z', simp add: nat add le Zero_int_def) | 
| 23164 | 317 | |
| 318 | lemma nat_diff_distrib: | |
| 319 | "[| (0::int) \<le> z'; z' \<le> z |] ==> nat (z-z') = nat z - nat z'" | |
| 320 | by (cases z, cases z', | |
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 321 | simp add: nat add minus diff_minus le Zero_int_def) | 
| 23164 | 322 | |
| 24196 | 323 | lemma nat_zminus_int [simp]: "nat (- (of_nat n)) = 0" | 
| 23365 | 324 | by (simp add: int_def minus nat Zero_int_def) | 
| 23164 | 325 | |
| 24196 | 326 | lemma zless_nat_eq_int_zless: "(m < nat z) = (of_nat m < z)" | 
| 23365 | 327 | by (cases z, simp add: nat less int_def, arith) | 
| 23164 | 328 | |
| 329 | ||
| 24196 | 330 | subsection{*Lemmas about the Function @{term of_nat} and Orderings*}
 | 
| 23164 | 331 | |
| 24196 | 332 | lemma negative_zless_0: "- (of_nat (Suc n)) < (0 \<Colon> int)" | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 333 | by (simp add: order_less_le del: of_nat_Suc) | 
| 23164 | 334 | |
| 24196 | 335 | lemma negative_zless [iff]: "- (of_nat (Suc n)) < (of_nat m \<Colon> int)" | 
| 23365 | 336 | by (rule negative_zless_0 [THEN order_less_le_trans], simp) | 
| 23164 | 337 | |
| 24196 | 338 | lemma negative_zle_0: "- of_nat n \<le> (0 \<Colon> int)" | 
| 23164 | 339 | by (simp add: minus_le_iff) | 
| 340 | ||
| 24196 | 341 | lemma negative_zle [iff]: "- of_nat n \<le> (of_nat m \<Colon> int)" | 
| 23365 | 342 | by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff]) | 
| 23164 | 343 | |
| 24196 | 344 | lemma not_zle_0_negative [simp]: "~ (0 \<le> - (of_nat (Suc n) \<Colon> int))" | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 345 | by (subst le_minus_iff, simp del: of_nat_Suc) | 
| 23164 | 346 | |
| 24196 | 347 | lemma int_zle_neg: "((of_nat n \<Colon> int) \<le> - of_nat m) = (n = 0 & m = 0)" | 
| 23365 | 348 | by (simp add: int_def le minus Zero_int_def) | 
| 23164 | 349 | |
| 24196 | 350 | lemma not_int_zless_negative [simp]: "~ ((of_nat n \<Colon> int) < - of_nat m)" | 
| 23164 | 351 | by (simp add: linorder_not_less) | 
| 352 | ||
| 24196 | 353 | lemma negative_eq_positive [simp]: "((- of_nat n \<Colon> int) = of_nat m) = (n = 0 & m = 0)" | 
| 354 | by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg) | |
| 23164 | 355 | |
| 24196 | 356 | lemma zle_iff_zadd: "(w\<Colon>int) \<le> z \<longleftrightarrow> (\<exists>n. z = w + of_nat n)" | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 357 | proof - | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 358 | have "(w \<le> z) = (0 \<le> z - w)" | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 359 | by (simp only: le_diff_eq add_0_left) | 
| 24196 | 360 | also have "\<dots> = (\<exists>n. z - w = of_nat n)" | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 361 | by (auto elim: zero_le_imp_eq_int) | 
| 24196 | 362 | also have "\<dots> = (\<exists>n. z = w + of_nat n)" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23438diff
changeset | 363 | by (simp only: group_simps) | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 364 | finally show ?thesis . | 
| 23164 | 365 | qed | 
| 366 | ||
| 24196 | 367 | lemma zadd_int_left: "of_nat m + (of_nat n + z) = of_nat (m + n) + (z\<Colon>int)" | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 368 | by simp | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 369 | |
| 24196 | 370 | lemma int_Suc0_eq_1: "of_nat (Suc 0) = (1\<Colon>int)" | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 371 | by simp | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 372 | |
| 23164 | 373 | text{*This version is proved for all ordered rings, not just integers!
 | 
| 374 |       It is proved here because attribute @{text arith_split} is not available
 | |
| 375 |       in theory @{text Ring_and_Field}.
 | |
| 376 |       But is it really better than just rewriting with @{text abs_if}?*}
 | |
| 24286 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 paulson parents: 
24196diff
changeset | 377 | lemma abs_split [arith_split,noatp]: | 
| 23164 | 378 | "P(abs(a::'a::ordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))" | 
| 379 | by (force dest: order_less_le_trans simp add: abs_if linorder_not_less) | |
| 380 | ||
| 381 | ||
| 382 | subsection {* Constants @{term neg} and @{term iszero} *}
 | |
| 383 | ||
| 384 | definition | |
| 385 | neg :: "'a\<Colon>ordered_idom \<Rightarrow> bool" | |
| 386 | where | |
| 25164 | 387 | "neg Z \<longleftrightarrow> Z < 0" | 
| 23164 | 388 | |
| 389 | definition (*for simplifying equalities*) | |
| 23276 
a70934b63910
generalize of_nat and related constants to class semiring_1
 huffman parents: 
23164diff
changeset | 390 | iszero :: "'a\<Colon>semiring_1 \<Rightarrow> bool" | 
| 23164 | 391 | where | 
| 392 | "iszero z \<longleftrightarrow> z = 0" | |
| 393 | ||
| 24196 | 394 | lemma not_neg_int [simp]: "~ neg (of_nat n)" | 
| 23164 | 395 | by (simp add: neg_def) | 
| 396 | ||
| 24196 | 397 | lemma neg_zminus_int [simp]: "neg (- (of_nat (Suc n)))" | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 398 | by (simp add: neg_def neg_less_0_iff_less del: of_nat_Suc) | 
| 23164 | 399 | |
| 400 | lemmas neg_eq_less_0 = neg_def | |
| 401 | ||
| 402 | lemma not_neg_eq_ge_0: "(~neg x) = (0 \<le> x)" | |
| 403 | by (simp add: neg_def linorder_not_less) | |
| 404 | ||
| 405 | ||
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 406 | text{*To simplify inequalities when Numeral1 can get simplified to 1*}
 | 
| 23164 | 407 | |
| 408 | lemma not_neg_0: "~ neg 0" | |
| 409 | by (simp add: One_int_def neg_def) | |
| 410 | ||
| 411 | lemma not_neg_1: "~ neg 1" | |
| 412 | by (simp add: neg_def linorder_not_less zero_le_one) | |
| 413 | ||
| 414 | lemma iszero_0: "iszero 0" | |
| 415 | by (simp add: iszero_def) | |
| 416 | ||
| 417 | lemma not_iszero_1: "~ iszero 1" | |
| 418 | by (simp add: iszero_def eq_commute) | |
| 419 | ||
| 420 | lemma neg_nat: "neg z ==> nat z = 0" | |
| 421 | by (simp add: neg_def order_less_imp_le) | |
| 422 | ||
| 24196 | 423 | lemma not_neg_nat: "~ neg z ==> of_nat (nat z) = z" | 
| 23164 | 424 | by (simp add: linorder_not_less neg_def) | 
| 425 | ||
| 426 | ||
| 23852 | 427 | subsection{*Embedding of the Integers into any @{text ring_1}: @{term of_int}*}
 | 
| 23164 | 428 | |
| 25193 | 429 | context ring_1 | 
| 430 | begin | |
| 431 | ||
| 25230 | 432 | term of_nat | 
| 433 | ||
| 23950 | 434 | definition | 
| 25193 | 435 | of_int :: "int \<Rightarrow> 'a" | 
| 23950 | 436 | where | 
| 437 |   "of_int z = contents (\<Union>(i, j) \<in> Rep_Integ z. { of_nat i - of_nat j })"
 | |
| 23852 | 438 | lemmas [code func del] = of_int_def | 
| 23164 | 439 | |
| 440 | lemma of_int: "of_int (Abs_Integ (intrel `` {(i,j)})) = of_nat i - of_nat j"
 | |
| 441 | proof - | |
| 442 |   have "(\<lambda>(i,j). { of_nat i - (of_nat j :: 'a) }) respects intrel"
 | |
| 443 | by (simp add: congruent_def compare_rls of_nat_add [symmetric] | |
| 444 | del: of_nat_add) | |
| 445 | thus ?thesis | |
| 446 | by (simp add: of_int_def UN_equiv_class [OF equiv_intrel]) | |
| 447 | qed | |
| 448 | ||
| 449 | lemma of_int_0 [simp]: "of_int 0 = 0" | |
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 450 | by (simp add: of_int Zero_int_def) | 
| 23164 | 451 | |
| 452 | lemma of_int_1 [simp]: "of_int 1 = 1" | |
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 453 | by (simp add: of_int One_int_def) | 
| 23164 | 454 | |
| 455 | lemma of_int_add [simp]: "of_int (w+z) = of_int w + of_int z" | |
| 456 | by (cases w, cases z, simp add: compare_rls of_int add) | |
| 457 | ||
| 458 | lemma of_int_minus [simp]: "of_int (-z) = - (of_int z)" | |
| 459 | by (cases z, simp add: compare_rls of_int minus) | |
| 460 | ||
| 461 | lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z" | |
| 462 | apply (cases w, cases z) | |
| 463 | apply (simp add: compare_rls of_int left_diff_distrib right_diff_distrib | |
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 464 | mult add_ac of_nat_mult) | 
| 23164 | 465 | done | 
| 466 | ||
| 25193 | 467 | text{*Collapse nested embeddings*}
 | 
| 468 | lemma of_int_of_nat_eq [simp]: "of_int (Nat.of_nat n) = of_nat n" | |
| 469 | by (induct n, auto) | |
| 470 | ||
| 471 | end | |
| 472 | ||
| 473 | lemma of_int_diff [simp]: "of_int (w-z) = of_int w - of_int z" | |
| 474 | by (simp add: diff_minus) | |
| 475 | ||
| 23164 | 476 | lemma of_int_le_iff [simp]: | 
| 477 | "(of_int w \<le> (of_int z::'a::ordered_idom)) = (w \<le> z)" | |
| 478 | apply (cases w) | |
| 479 | apply (cases z) | |
| 480 | apply (simp add: compare_rls of_int le diff_int_def add minus | |
| 481 | of_nat_add [symmetric] del: of_nat_add) | |
| 482 | done | |
| 483 | ||
| 484 | text{*Special cases where either operand is zero*}
 | |
| 485 | lemmas of_int_0_le_iff [simp] = of_int_le_iff [of 0, simplified] | |
| 486 | lemmas of_int_le_0_iff [simp] = of_int_le_iff [of _ 0, simplified] | |
| 487 | ||
| 488 | lemma of_int_less_iff [simp]: | |
| 489 | "(of_int w < (of_int z::'a::ordered_idom)) = (w < z)" | |
| 490 | by (simp add: linorder_not_le [symmetric]) | |
| 491 | ||
| 492 | text{*Special cases where either operand is zero*}
 | |
| 493 | lemmas of_int_0_less_iff [simp] = of_int_less_iff [of 0, simplified] | |
| 494 | lemmas of_int_less_0_iff [simp] = of_int_less_iff [of _ 0, simplified] | |
| 495 | ||
| 496 | text{*Class for unital rings with characteristic zero.
 | |
| 497 | Includes non-ordered rings like the complex numbers.*} | |
| 23950 | 498 | class ring_char_0 = ring_1 + semiring_char_0 | 
| 25193 | 499 | begin | 
| 23164 | 500 | |
| 501 | lemma of_int_eq_iff [simp]: | |
| 25193 | 502 | "of_int w = of_int z \<longleftrightarrow> w = z" | 
| 23282 
dfc459989d24
add axclass semiring_char_0 for types where of_nat is injective
 huffman parents: 
23276diff
changeset | 503 | apply (cases w, cases z, simp add: of_int) | 
| 
dfc459989d24
add axclass semiring_char_0 for types where of_nat is injective
 huffman parents: 
23276diff
changeset | 504 | apply (simp only: diff_eq_eq diff_add_eq eq_diff_eq) | 
| 
dfc459989d24
add axclass semiring_char_0 for types where of_nat is injective
 huffman parents: 
23276diff
changeset | 505 | apply (simp only: of_nat_add [symmetric] of_nat_eq_iff) | 
| 
dfc459989d24
add axclass semiring_char_0 for types where of_nat is injective
 huffman parents: 
23276diff
changeset | 506 | done | 
| 23164 | 507 | |
| 508 | text{*Special cases where either operand is zero*}
 | |
| 509 | lemmas of_int_0_eq_iff [simp] = of_int_eq_iff [of 0, simplified] | |
| 510 | lemmas of_int_eq_0_iff [simp] = of_int_eq_iff [of _ 0, simplified] | |
| 511 | ||
| 25193 | 512 | end | 
| 513 | ||
| 514 | text{*Every @{text ordered_idom} has characteristic zero.*}
 | |
| 515 | instance ordered_idom \<subseteq> ring_char_0 .. | |
| 516 | ||
| 517 | lemma of_int_eq_id [simp]: "of_int = id" | |
| 23164 | 518 | proof | 
| 25193 | 519 | fix z show "of_int z = id z" | 
| 520 | by (cases z) (simp add: of_int add minus int_def diff_minus) | |
| 23164 | 521 | qed | 
| 522 | ||
| 25230 | 523 | context ring_1 | 
| 524 | begin | |
| 525 | ||
| 25193 | 526 | lemma of_nat_nat: "0 \<le> z \<Longrightarrow> of_nat (nat z) = of_int z" | 
| 25230 | 527 | by (cases z rule: eq_Abs_Integ) | 
| 23438 
dd824e86fa8a
remove simp attribute from of_nat_diff, for backward compatibility with zdiff_int
 huffman parents: 
23431diff
changeset | 528 | (simp add: nat le of_int Zero_int_def of_nat_diff) | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 529 | |
| 25230 | 530 | end | 
| 531 | ||
| 23164 | 532 | |
| 533 | subsection{*The Set of Integers*}
 | |
| 534 | ||
| 25193 | 535 | context ring_1 | 
| 536 | begin | |
| 537 | ||
| 538 | definition | |
| 539 | Ints :: "'a set" | |
| 540 | where | |
| 541 | "Ints = range of_int" | |
| 542 | ||
| 543 | end | |
| 23164 | 544 | |
| 545 | notation (xsymbols) | |
| 546 |   Ints  ("\<int>")
 | |
| 547 | ||
| 25193 | 548 | context ring_1 | 
| 549 | begin | |
| 550 | ||
| 551 | lemma Ints_0 [simp]: "0 \<in> \<int>" | |
| 23164 | 552 | apply (simp add: Ints_def) | 
| 553 | apply (rule range_eqI) | |
| 554 | apply (rule of_int_0 [symmetric]) | |
| 555 | done | |
| 556 | ||
| 25193 | 557 | lemma Ints_1 [simp]: "1 \<in> \<int>" | 
| 23164 | 558 | apply (simp add: Ints_def) | 
| 559 | apply (rule range_eqI) | |
| 560 | apply (rule of_int_1 [symmetric]) | |
| 561 | done | |
| 562 | ||
| 25193 | 563 | lemma Ints_add [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a + b \<in> \<int>" | 
| 23164 | 564 | apply (auto simp add: Ints_def) | 
| 565 | apply (rule range_eqI) | |
| 566 | apply (rule of_int_add [symmetric]) | |
| 567 | done | |
| 568 | ||
| 25193 | 569 | lemma Ints_minus [simp]: "a \<in> \<int> \<Longrightarrow> -a \<in> \<int>" | 
| 23164 | 570 | apply (auto simp add: Ints_def) | 
| 571 | apply (rule range_eqI) | |
| 572 | apply (rule of_int_minus [symmetric]) | |
| 573 | done | |
| 574 | ||
| 25193 | 575 | lemma Ints_mult [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a * b \<in> \<int>" | 
| 23164 | 576 | apply (auto simp add: Ints_def) | 
| 577 | apply (rule range_eqI) | |
| 578 | apply (rule of_int_mult [symmetric]) | |
| 579 | done | |
| 580 | ||
| 581 | lemma Ints_cases [cases set: Ints]: | |
| 582 | assumes "q \<in> \<int>" | |
| 583 | obtains (of_int) z where "q = of_int z" | |
| 584 | unfolding Ints_def | |
| 585 | proof - | |
| 586 | from `q \<in> \<int>` have "q \<in> range of_int" unfolding Ints_def . | |
| 587 | then obtain z where "q = of_int z" .. | |
| 588 | then show thesis .. | |
| 589 | qed | |
| 590 | ||
| 591 | lemma Ints_induct [case_names of_int, induct set: Ints]: | |
| 25193 | 592 | "q \<in> \<int> \<Longrightarrow> (\<And>z. P (of_int z)) \<Longrightarrow> P q" | 
| 23164 | 593 | by (rule Ints_cases) auto | 
| 594 | ||
| 25193 | 595 | end | 
| 596 | ||
| 597 | lemma Ints_diff [simp]: "a \<in> \<int> \<Longrightarrow> b \<in> \<int> \<Longrightarrow> a-b \<in> \<int>" | |
| 598 | apply (auto simp add: Ints_def) | |
| 599 | apply (rule range_eqI) | |
| 600 | apply (rule of_int_diff [symmetric]) | |
| 601 | done | |
| 602 | ||
| 23164 | 603 | |
| 24728 | 604 | subsection {* @{term setsum} and @{term setprod} *}
 | 
| 605 | ||
| 606 | text {*By Jeremy Avigad*}
 | |
| 607 | ||
| 608 | lemma of_nat_setsum: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))" | |
| 609 | apply (cases "finite A") | |
| 610 | apply (erule finite_induct, auto) | |
| 611 | done | |
| 612 | ||
| 613 | lemma of_int_setsum: "of_int (setsum f A) = (\<Sum>x\<in>A. of_int(f x))" | |
| 614 | apply (cases "finite A") | |
| 615 | apply (erule finite_induct, auto) | |
| 616 | done | |
| 617 | ||
| 618 | lemma of_nat_setprod: "of_nat (setprod f A) = (\<Prod>x\<in>A. of_nat(f x))" | |
| 619 | apply (cases "finite A") | |
| 620 | apply (erule finite_induct, auto simp add: of_nat_mult) | |
| 621 | done | |
| 622 | ||
| 623 | lemma of_int_setprod: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))" | |
| 624 | apply (cases "finite A") | |
| 625 | apply (erule finite_induct, auto) | |
| 626 | done | |
| 627 | ||
| 628 | lemma setprod_nonzero_nat: | |
| 629 | "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::nat)) ==> setprod f A \<noteq> 0" | |
| 630 | by (rule setprod_nonzero, auto) | |
| 631 | ||
| 632 | lemma setprod_zero_eq_nat: | |
| 633 | "finite A ==> (setprod f A = (0::nat)) = (\<exists>x \<in> A. f x = 0)" | |
| 634 | by (rule setprod_zero_eq, auto) | |
| 635 | ||
| 636 | lemma setprod_nonzero_int: | |
| 637 | "finite A ==> (\<forall>x \<in> A. f x \<noteq> (0::int)) ==> setprod f A \<noteq> 0" | |
| 638 | by (rule setprod_nonzero, auto) | |
| 639 | ||
| 640 | lemma setprod_zero_eq_int: | |
| 641 | "finite A ==> (setprod f A = (0::int)) = (\<exists>x \<in> A. f x = 0)" | |
| 642 | by (rule setprod_zero_eq, auto) | |
| 643 | ||
| 644 | lemmas int_setsum = of_nat_setsum [where 'a=int] | |
| 645 | lemmas int_setprod = of_nat_setprod [where 'a=int] | |
| 646 | ||
| 647 | ||
| 23164 | 648 | subsection {* Further properties *}
 | 
| 649 | ||
| 650 | text{*Now we replace the case analysis rule by a more conventional one:
 | |
| 651 | whether an integer is negative or not.*} | |
| 652 | ||
| 23365 | 653 | lemma zless_iff_Suc_zadd: | 
| 24196 | 654 | "(w \<Colon> int) < z \<longleftrightarrow> (\<exists>n. z = w + of_nat (Suc n))" | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 655 | apply (cases z, cases w) | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 656 | apply (auto simp add: less add int_def) | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 657 | apply (rename_tac a b c d) | 
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 658 | apply (rule_tac x="a+d - Suc(c+b)" in exI) | 
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 659 | apply arith | 
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 660 | done | 
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 661 | |
| 24196 | 662 | lemma negD: "(x \<Colon> int) < 0 \<Longrightarrow> \<exists>n. x = - (of_nat (Suc n))" | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 663 | apply (cases x) | 
| 23365 | 664 | apply (auto simp add: le minus Zero_int_def int_def order_less_le) | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 665 | apply (rule_tac x="y - Suc x" in exI, arith) | 
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 666 | done | 
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 667 | |
| 23365 | 668 | theorem int_cases [cases type: int, case_names nonneg neg]: | 
| 24196 | 669 | "[|!! n. (z \<Colon> int) = of_nat n ==> P; !! n. z = - (of_nat (Suc n)) ==> P |] ==> P" | 
| 23365 | 670 | apply (cases "z < 0", blast dest!: negD) | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 671 | apply (simp add: linorder_not_less del: of_nat_Suc) | 
| 23365 | 672 | apply (blast dest: nat_0_le [THEN sym]) | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 673 | done | 
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 674 | |
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 675 | theorem int_induct [induct type: int, case_names nonneg neg]: | 
| 24196 | 676 | "[|!! n. P (of_nat n \<Colon> int); !!n. P (- (of_nat (Suc n))) |] ==> P z" | 
| 23365 | 677 | by (cases z rule: int_cases) auto | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 678 | |
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 679 | text{*Contributed by Brian Huffman*}
 | 
| 23365 | 680 | theorem int_diff_cases [case_names diff]: | 
| 24196 | 681 | assumes prem: "!!m n. (z\<Colon>int) = of_nat m - of_nat n ==> P" shows "P" | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 682 | apply (cases z rule: eq_Abs_Integ) | 
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 683 | apply (rule_tac m=x and n=y in prem) | 
| 23365 | 684 | apply (simp add: int_def diff_def minus add) | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 685 | done | 
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 686 | |
| 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 687 | |
| 23365 | 688 | subsection {* Legacy theorems *}
 | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 689 | |
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 690 | lemmas zminus_zminus = minus_minus [of "?z::int"] | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 691 | lemmas zminus_0 = minus_zero [where 'a=int] | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 692 | lemmas zminus_zadd_distrib = minus_add_distrib [of "?z::int" "?w"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 693 | lemmas zadd_commute = add_commute [of "?z::int" "?w"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 694 | lemmas zadd_assoc = add_assoc [of "?z1.0::int" "?z2.0" "?z3.0"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 695 | lemmas zadd_left_commute = add_left_commute [of "?x::int" "?y" "?z"] | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 696 | lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 697 | lemmas zmult_ac = OrderedGroup.mult_ac | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 698 | lemmas zadd_0 = OrderedGroup.add_0_left [of "?z::int"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 699 | lemmas zadd_0_right = OrderedGroup.add_0_left [of "?z::int"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 700 | lemmas zadd_zminus_inverse2 = left_minus [of "?z::int"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 701 | lemmas zmult_zminus = mult_minus_left [of "?z::int" "?w"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 702 | lemmas zmult_commute = mult_commute [of "?z::int" "?w"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 703 | lemmas zmult_assoc = mult_assoc [of "?z1.0::int" "?z2.0" "?z3.0"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 704 | lemmas zadd_zmult_distrib = left_distrib [of "?z1.0::int" "?z2.0" "?w"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 705 | lemmas zadd_zmult_distrib2 = right_distrib [of "?w::int" "?z1.0" "?z2.0"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 706 | lemmas zdiff_zmult_distrib = left_diff_distrib [of "?z1.0::int" "?z2.0" "?w"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 707 | lemmas zdiff_zmult_distrib2 = right_diff_distrib [of "?w::int" "?z1.0" "?z2.0"] | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 708 | |
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 709 | lemmas int_distrib = | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 710 | zadd_zmult_distrib zadd_zmult_distrib2 | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 711 | zdiff_zmult_distrib zdiff_zmult_distrib2 | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 712 | |
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 713 | lemmas zmult_1 = mult_1_left [of "?z::int"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 714 | lemmas zmult_1_right = mult_1_right [of "?z::int"] | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 715 | |
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 716 | lemmas zle_refl = order_refl [of "?w::int"] | 
| 23402 | 717 | lemmas zle_trans = order_trans [where 'a=int and x="?i" and y="?j" and z="?k"] | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 718 | lemmas zle_anti_sym = order_antisym [of "?z::int" "?w"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 719 | lemmas zle_linear = linorder_linear [of "?z::int" "?w"] | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 720 | lemmas zless_linear = linorder_less_linear [where 'a = int] | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 721 | |
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 722 | lemmas zadd_left_mono = add_left_mono [of "?i::int" "?j" "?k"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 723 | lemmas zadd_strict_right_mono = add_strict_right_mono [of "?i::int" "?j" "?k"] | 
| 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 724 | lemmas zadd_zless_mono = add_less_le_mono [of "?w'::int" "?w" "?z'" "?z"] | 
| 23372 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 725 | |
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 726 | lemmas int_0_less_1 = zero_less_one [where 'a=int] | 
| 
0035be079bee
clean up instance proofs; reorganize section headings
 huffman parents: 
23365diff
changeset | 727 | lemmas int_0_neq_1 = zero_neq_one [where 'a=int] | 
| 23303 
6091e530ff77
add abbreviation int_of_nat for of_nat::nat=>int;
 huffman parents: 
23299diff
changeset | 728 | |
| 23365 | 729 | lemmas inj_int = inj_of_nat [where 'a=int] | 
| 730 | lemmas int_int_eq = of_nat_eq_iff [where 'a=int] | |
| 731 | lemmas zadd_int = of_nat_add [where 'a=int, symmetric] | |
| 732 | lemmas int_mult = of_nat_mult [where 'a=int] | |
| 733 | lemmas zmult_int = of_nat_mult [where 'a=int, symmetric] | |
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 734 | lemmas int_eq_0_conv = of_nat_eq_0_iff [where 'a=int and m="?n"] | 
| 23365 | 735 | lemmas zless_int = of_nat_less_iff [where 'a=int] | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 736 | lemmas int_less_0_conv = of_nat_less_0_iff [where 'a=int and m="?k"] | 
| 23365 | 737 | lemmas zero_less_int_conv = of_nat_0_less_iff [where 'a=int] | 
| 738 | lemmas zle_int = of_nat_le_iff [where 'a=int] | |
| 739 | lemmas zero_zle_int = of_nat_0_le_iff [where 'a=int] | |
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 740 | lemmas int_le_0_conv = of_nat_le_0_iff [where 'a=int and m="?n"] | 
| 24196 | 741 | lemmas int_0 = of_nat_0 [where 'a=int] | 
| 23365 | 742 | lemmas int_1 = of_nat_1 [where 'a=int] | 
| 24196 | 743 | lemmas int_Suc = of_nat_Suc [where 'a=int] | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23402diff
changeset | 744 | lemmas abs_int_eq = abs_of_nat [where 'a=int and n="?m"] | 
| 23365 | 745 | lemmas of_int_int_eq = of_int_of_nat_eq [where 'a=int] | 
| 746 | lemmas zdiff_int = of_nat_diff [where 'a=int, symmetric] | |
| 747 | lemmas zless_le = less_int_def [THEN meta_eq_to_obj_eq] | |
| 748 | lemmas int_eq_of_nat = TrueI | |
| 23164 | 749 | |
| 23365 | 750 | abbreviation | 
| 24196 | 751 | int :: "nat \<Rightarrow> int" | 
| 752 | where | |
| 753 | "int \<equiv> of_nat" | |
| 754 | ||
| 23164 | 755 | end |