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