author | paulson |
Fri, 09 Jan 2004 10:46:18 +0100 | |
changeset 14348 | 744c868ee0b7 |
parent 14341 | a09441bd4f1e |
child 14378 | 69c4d5997669 |
permissions | -rw-r--r-- |
5508 | 1 |
(* Title: IntDef.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
6 |
*) |
|
7 |
||
14271 | 8 |
header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*} |
9 |
||
14259 | 10 |
theory IntDef = Equiv + NatArith: |
5508 | 11 |
constdefs |
14271 | 12 |
intrel :: "((nat * nat) * (nat * nat)) set" |
13 |
"intrel == {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}" |
|
5508 | 14 |
|
15 |
typedef (Integ) |
|
14259 | 16 |
int = "UNIV//intrel" |
17 |
by (auto simp add: quotient_def) |
|
5508 | 18 |
|
14259 | 19 |
instance int :: ord .. |
20 |
instance int :: zero .. |
|
21 |
instance int :: one .. |
|
22 |
instance int :: plus .. |
|
23 |
instance int :: times .. |
|
24 |
instance int :: minus .. |
|
5508 | 25 |
|
26 |
constdefs |
|
27 |
||
14259 | 28 |
int :: "nat => int" |
10834 | 29 |
"int m == Abs_Integ(intrel `` {(m,0)})" |
5508 | 30 |
|
14259 | 31 |
neg :: "int => bool" |
14271 | 32 |
"neg(Z) == \<exists>x y. x<y & (x,y::nat):Rep_Integ(Z)" |
5508 | 33 |
|
34 |
(*For simplifying equalities*) |
|
14259 | 35 |
iszero :: "int => bool" |
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11713
diff
changeset
|
36 |
"iszero z == z = (0::int)" |
5508 | 37 |
|
14269 | 38 |
defs (overloaded) |
9392 | 39 |
|
14271 | 40 |
zminus_def: "- Z == Abs_Integ(\<Union>(x,y) \<in> Rep_Integ(Z). intrel``{(y,x)})" |
41 |
||
14259 | 42 |
Zero_int_def: "0 == int 0" |
14271 | 43 |
One_int_def: "1 == int 1" |
8937 | 44 |
|
14259 | 45 |
zadd_def: |
5594 | 46 |
"z + w == |
14271 | 47 |
Abs_Integ(\<Union>(x1,y1) \<in> Rep_Integ(z). \<Union>(x2,y2) \<in> Rep_Integ(w). |
10834 | 48 |
intrel``{(x1+x2, y1+y2)})" |
5508 | 49 |
|
14259 | 50 |
zdiff_def: "z - (w::int) == z + (-w)" |
5508 | 51 |
|
14259 | 52 |
zless_def: "z<w == neg(z - w)" |
5508 | 53 |
|
14259 | 54 |
zle_def: "z <= (w::int) == ~(w < z)" |
5508 | 55 |
|
14259 | 56 |
zmult_def: |
5594 | 57 |
"z * w == |
14271 | 58 |
Abs_Integ(\<Union>(x1,y1) \<in> Rep_Integ(z). \<Union>(x2,y2) \<in> Rep_Integ(w). |
10834 | 59 |
intrel``{(x1*x2 + y1*y2, x1*y2 + y1*x2)})" |
5508 | 60 |
|
14271 | 61 |
lemma intrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> intrel) = (x1+y2 = x2+y1)" |
14259 | 62 |
by (unfold intrel_def, blast) |
63 |
||
64 |
lemma equiv_intrel: "equiv UNIV intrel" |
|
65 |
by (unfold intrel_def equiv_def refl_def sym_def trans_def, auto) |
|
66 |
||
67 |
lemmas equiv_intrel_iff = |
|
68 |
eq_equiv_class_iff [OF equiv_intrel UNIV_I UNIV_I, simp] |
|
69 |
||
70 |
lemma intrel_in_integ [simp]: "intrel``{(x,y)}:Integ" |
|
71 |
by (unfold Integ_def intrel_def quotient_def, fast) |
|
72 |
||
73 |
lemma inj_on_Abs_Integ: "inj_on Abs_Integ Integ" |
|
74 |
apply (rule inj_on_inverseI) |
|
75 |
apply (erule Abs_Integ_inverse) |
|
76 |
done |
|
77 |
||
78 |
declare inj_on_Abs_Integ [THEN inj_on_iff, simp] |
|
79 |
Abs_Integ_inverse [simp] |
|
80 |
||
81 |
lemma inj_Rep_Integ: "inj(Rep_Integ)" |
|
82 |
apply (rule inj_on_inverseI) |
|
83 |
apply (rule Rep_Integ_inverse) |
|
84 |
done |
|
85 |
||
86 |
||
87 |
(** int: the injection from "nat" to "int" **) |
|
88 |
||
89 |
lemma inj_int: "inj int" |
|
90 |
apply (rule inj_onI) |
|
91 |
apply (unfold int_def) |
|
92 |
apply (drule inj_on_Abs_Integ [THEN inj_onD]) |
|
93 |
apply (rule intrel_in_integ)+ |
|
94 |
apply (drule eq_equiv_class) |
|
95 |
apply (rule equiv_intrel, fast) |
|
96 |
apply (simp add: intrel_def) |
|
97 |
done |
|
98 |
||
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
99 |
lemma int_int_eq [iff]: "(int m = int n) = (m = n)" |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
100 |
by (fast elim!: inj_int [THEN injD]) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
101 |
|
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
102 |
|
14259 | 103 |
|
104 |
subsection{*zminus: unary negation on Integ*} |
|
105 |
||
106 |
lemma zminus_congruent: "congruent intrel (%(x,y). intrel``{(y,x)})" |
|
107 |
apply (unfold congruent_def intrel_def) |
|
108 |
apply (auto simp add: add_ac) |
|
109 |
done |
|
110 |
||
111 |
lemma zminus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})" |
|
112 |
by (simp add: zminus_def equiv_intrel [THEN UN_equiv_class] zminus_congruent) |
|
113 |
||
114 |
(*Every integer can be written in the form Abs_Integ(...) *) |
|
115 |
lemma eq_Abs_Integ: |
|
116 |
"(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P" |
|
117 |
apply (rule_tac x1=z in Rep_Integ [unfolded Integ_def, THEN quotientE]) |
|
118 |
apply (drule_tac f = Abs_Integ in arg_cong) |
|
119 |
apply (rule_tac p = x in PairE) |
|
120 |
apply (simp add: Rep_Integ_inverse) |
|
121 |
done |
|
122 |
||
123 |
lemma zminus_zminus [simp]: "- (- z) = (z::int)" |
|
124 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
125 |
apply (simp (no_asm_simp) add: zminus) |
|
126 |
done |
|
127 |
||
128 |
lemma inj_zminus: "inj(%z::int. -z)" |
|
129 |
apply (rule inj_onI) |
|
130 |
apply (drule_tac f = uminus in arg_cong, simp) |
|
131 |
done |
|
132 |
||
133 |
lemma zminus_0 [simp]: "- 0 = (0::int)" |
|
134 |
by (simp add: int_def Zero_int_def zminus) |
|
135 |
||
136 |
||
137 |
subsection{*neg: the test for negative integers*} |
|
138 |
||
139 |
||
140 |
lemma not_neg_int [simp]: "~ neg(int n)" |
|
141 |
by (simp add: neg_def int_def) |
|
142 |
||
143 |
lemma neg_zminus_int [simp]: "neg(- (int (Suc n)))" |
|
144 |
by (simp add: neg_def int_def zminus) |
|
145 |
||
146 |
||
147 |
subsection{*zadd: addition on Integ*} |
|
148 |
||
149 |
lemma zadd: |
|
150 |
"Abs_Integ(intrel``{(x1,y1)}) + Abs_Integ(intrel``{(x2,y2)}) = |
|
151 |
Abs_Integ(intrel``{(x1+x2, y1+y2)})" |
|
152 |
apply (simp add: zadd_def UN_UN_split_split_eq) |
|
153 |
apply (subst equiv_intrel [THEN UN_equiv_class2]) |
|
154 |
apply (auto simp add: congruent2_def) |
|
155 |
done |
|
156 |
||
157 |
lemma zminus_zadd_distrib [simp]: "- (z + w) = (- z) + (- w::int)" |
|
158 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
159 |
apply (rule_tac z = w in eq_Abs_Integ) |
|
160 |
apply (simp (no_asm_simp) add: zminus zadd) |
|
161 |
done |
|
162 |
||
163 |
lemma zadd_commute: "(z::int) + w = w + z" |
|
164 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
165 |
apply (rule_tac z = w in eq_Abs_Integ) |
|
166 |
apply (simp (no_asm_simp) add: add_ac zadd) |
|
167 |
done |
|
168 |
||
169 |
lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)" |
|
170 |
apply (rule_tac z = z1 in eq_Abs_Integ) |
|
171 |
apply (rule_tac z = z2 in eq_Abs_Integ) |
|
172 |
apply (rule_tac z = z3 in eq_Abs_Integ) |
|
173 |
apply (simp (no_asm_simp) add: zadd add_assoc) |
|
174 |
done |
|
175 |
||
176 |
(*For AC rewriting*) |
|
14271 | 177 |
lemma zadd_left_commute: "x + (y + z) = y + ((x + z) ::int)" |
14259 | 178 |
apply (rule mk_left_commute [of "op +"]) |
179 |
apply (rule zadd_assoc) |
|
180 |
apply (rule zadd_commute) |
|
181 |
done |
|
182 |
||
183 |
(*Integer addition is an AC operator*) |
|
184 |
lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute |
|
185 |
||
14271 | 186 |
lemmas zmult_ac = Ring_and_Field.mult_ac |
187 |
||
14259 | 188 |
lemma zadd_int: "(int m) + (int n) = int (m + n)" |
189 |
by (simp add: int_def zadd) |
|
190 |
||
191 |
lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z" |
|
192 |
by (simp add: zadd_int zadd_assoc [symmetric]) |
|
193 |
||
194 |
lemma int_Suc: "int (Suc m) = 1 + (int m)" |
|
195 |
by (simp add: One_int_def zadd_int) |
|
196 |
||
197 |
(*also for the instance declaration int :: plus_ac0*) |
|
198 |
lemma zadd_0 [simp]: "(0::int) + z = z" |
|
199 |
apply (unfold Zero_int_def int_def) |
|
200 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
201 |
apply (simp (no_asm_simp) add: zadd) |
|
202 |
done |
|
203 |
||
204 |
lemma zadd_0_right [simp]: "z + (0::int) = z" |
|
205 |
by (rule trans [OF zadd_commute zadd_0]) |
|
206 |
||
207 |
lemma zadd_zminus_inverse [simp]: "z + (- z) = (0::int)" |
|
208 |
apply (unfold int_def Zero_int_def) |
|
209 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
210 |
apply (simp (no_asm_simp) add: zminus zadd add_commute) |
|
211 |
done |
|
212 |
||
213 |
lemma zadd_zminus_inverse2 [simp]: "(- z) + z = (0::int)" |
|
214 |
apply (rule zadd_commute [THEN trans]) |
|
215 |
apply (rule zadd_zminus_inverse) |
|
216 |
done |
|
217 |
||
218 |
lemma zadd_zminus_cancel [simp]: "z + (- z + w) = (w::int)" |
|
219 |
by (simp add: zadd_assoc [symmetric] zadd_0) |
|
220 |
||
221 |
lemma zminus_zadd_cancel [simp]: "(-z) + (z + w) = (w::int)" |
|
222 |
by (simp add: zadd_assoc [symmetric] zadd_0) |
|
223 |
||
224 |
lemma zdiff0 [simp]: "(0::int) - x = -x" |
|
225 |
by (simp add: zdiff_def) |
|
226 |
||
227 |
lemma zdiff0_right [simp]: "x - (0::int) = x" |
|
228 |
by (simp add: zdiff_def) |
|
229 |
||
230 |
lemma zdiff_self [simp]: "x - x = (0::int)" |
|
231 |
by (simp add: zdiff_def Zero_int_def) |
|
232 |
||
233 |
||
234 |
(** Lemmas **) |
|
235 |
||
236 |
lemma zadd_assoc_cong: "(z::int) + v = z' + v' ==> z + (v + w) = z' + (v' + w)" |
|
237 |
by (simp add: zadd_assoc [symmetric]) |
|
238 |
||
239 |
lemma zadd_assoc_swap: "(z::int) + (v + w) = v + (z + w)" |
|
240 |
by (rule zadd_commute [THEN zadd_assoc_cong]) |
|
241 |
||
242 |
||
243 |
subsection{*zmult: multiplication on Integ*} |
|
244 |
||
245 |
(*Congruence property for multiplication*) |
|
246 |
lemma zmult_congruent2: "congruent2 intrel |
|
247 |
(%p1 p2. (%(x1,y1). (%(x2,y2). |
|
248 |
intrel``{(x1*x2 + y1*y2, x1*y2 + y1*x2)}) p2) p1)" |
|
249 |
apply (rule equiv_intrel [THEN congruent2_commuteI]) |
|
250 |
apply (rule_tac [2] p=w in PairE) |
|
251 |
apply (force simp add: add_ac mult_ac, clarify) |
|
252 |
apply (simp (no_asm_simp) del: equiv_intrel_iff add: add_ac mult_ac) |
|
253 |
apply (rename_tac x1 x2 y1 y2 z1 z2) |
|
254 |
apply (rule equiv_class_eq [OF equiv_intrel intrel_iff [THEN iffD2]]) |
|
255 |
apply (simp add: intrel_def) |
|
256 |
apply (subgoal_tac "x1*z1 + y2*z1 = y1*z1 + x2*z1 & x1*z2 + y2*z2 = y1*z2 + x2*z2", arith) |
|
257 |
apply (simp add: add_mult_distrib [symmetric]) |
|
258 |
done |
|
259 |
||
260 |
lemma zmult: |
|
261 |
"Abs_Integ((intrel``{(x1,y1)})) * Abs_Integ((intrel``{(x2,y2)})) = |
|
262 |
Abs_Integ(intrel `` {(x1*x2 + y1*y2, x1*y2 + y1*x2)})" |
|
263 |
apply (unfold zmult_def) |
|
264 |
apply (simp (no_asm_simp) add: UN_UN_split_split_eq zmult_congruent2 equiv_intrel [THEN UN_equiv_class2]) |
|
265 |
done |
|
266 |
||
267 |
lemma zmult_zminus: "(- z) * w = - (z * (w::int))" |
|
268 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
269 |
apply (rule_tac z = w in eq_Abs_Integ) |
|
270 |
apply (simp (no_asm_simp) add: zminus zmult add_ac) |
|
271 |
done |
|
272 |
||
273 |
lemma zmult_commute: "(z::int) * w = w * z" |
|
274 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
275 |
apply (rule_tac z = w in eq_Abs_Integ) |
|
276 |
apply (simp (no_asm_simp) add: zmult add_ac mult_ac) |
|
277 |
done |
|
278 |
||
279 |
lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)" |
|
280 |
apply (rule_tac z = z1 in eq_Abs_Integ) |
|
281 |
apply (rule_tac z = z2 in eq_Abs_Integ) |
|
282 |
apply (rule_tac z = z3 in eq_Abs_Integ) |
|
283 |
apply (simp (no_asm_simp) add: add_mult_distrib2 zmult add_ac mult_ac) |
|
284 |
done |
|
285 |
||
286 |
lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)" |
|
287 |
apply (rule_tac z = z1 in eq_Abs_Integ) |
|
288 |
apply (rule_tac z = z2 in eq_Abs_Integ) |
|
289 |
apply (rule_tac z = w in eq_Abs_Integ) |
|
290 |
apply (simp add: add_mult_distrib2 zadd zmult add_ac mult_ac) |
|
291 |
done |
|
292 |
||
293 |
lemma zmult_zminus_right: "w * (- z) = - (w * (z::int))" |
|
294 |
by (simp add: zmult_commute [of w] zmult_zminus) |
|
295 |
||
296 |
lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)" |
|
297 |
by (simp add: zmult_commute [of w] zadd_zmult_distrib) |
|
298 |
||
299 |
lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)" |
|
300 |
apply (unfold zdiff_def) |
|
301 |
apply (subst zadd_zmult_distrib) |
|
302 |
apply (simp add: zmult_zminus) |
|
303 |
done |
|
304 |
||
305 |
lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)" |
|
306 |
by (simp add: zmult_commute [of w] zdiff_zmult_distrib) |
|
307 |
||
308 |
lemmas int_distrib = |
|
309 |
zadd_zmult_distrib zadd_zmult_distrib2 |
|
310 |
zdiff_zmult_distrib zdiff_zmult_distrib2 |
|
311 |
||
312 |
lemma zmult_int: "(int m) * (int n) = int (m * n)" |
|
313 |
by (simp add: int_def zmult) |
|
314 |
||
315 |
lemma zmult_0 [simp]: "0 * z = (0::int)" |
|
316 |
apply (unfold Zero_int_def int_def) |
|
317 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
318 |
apply (simp (no_asm_simp) add: zmult) |
|
319 |
done |
|
320 |
||
321 |
lemma zmult_1 [simp]: "(1::int) * z = z" |
|
322 |
apply (unfold One_int_def int_def) |
|
323 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
324 |
apply (simp (no_asm_simp) add: zmult) |
|
325 |
done |
|
326 |
||
327 |
lemma zmult_0_right [simp]: "z * 0 = (0::int)" |
|
328 |
by (rule trans [OF zmult_commute zmult_0]) |
|
329 |
||
330 |
lemma zmult_1_right [simp]: "z * (1::int) = z" |
|
331 |
by (rule trans [OF zmult_commute zmult_1]) |
|
332 |
||
333 |
||
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
334 |
text{*The Integers Form A Ring*} |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
335 |
instance int :: ring |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
336 |
proof |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
337 |
fix i j k :: int |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
338 |
show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
339 |
show "i + j = j + i" by (simp add: zadd_commute) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
340 |
show "0 + i = i" by simp |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
341 |
show "- i + i = 0" by simp |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
342 |
show "i - j = i + (-j)" by (simp add: zdiff_def) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
343 |
show "(i * j) * k = i * (j * k)" by (rule zmult_assoc) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
344 |
show "i * j = j * i" by (rule zmult_commute) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
345 |
show "1 * i = i" by simp |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
346 |
show "(i + j) * k = i * k + j * k" by (simp add: int_distrib) |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
347 |
show "0 \<noteq> (1::int)" |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14341
diff
changeset
|
348 |
by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq) |
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
349 |
assume eq: "k+i = k+j" |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
350 |
hence "(-k + k) + i = (-k + k) + j" by (simp only: eq zadd_assoc) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
351 |
thus "i = j" by simp |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
352 |
qed |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
353 |
|
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
354 |
|
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
355 |
subsection{*Theorems about the Ordering*} |
14259 | 356 |
|
357 |
(*This lemma allows direct proofs of other <-properties*) |
|
358 |
lemma zless_iff_Suc_zadd: |
|
14271 | 359 |
"(w < z) = (\<exists>n. z = w + int(Suc n))" |
14259 | 360 |
apply (unfold zless_def neg_def zdiff_def int_def) |
361 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
362 |
apply (rule_tac z = w in eq_Abs_Integ, clarify) |
|
363 |
apply (simp add: zadd zminus) |
|
364 |
apply (safe dest!: less_imp_Suc_add) |
|
365 |
apply (rule_tac x = k in exI) |
|
366 |
apply (simp_all add: add_ac) |
|
367 |
done |
|
368 |
||
369 |
lemma zless_zadd_Suc: "z < z + int (Suc n)" |
|
370 |
by (auto simp add: zless_iff_Suc_zadd zadd_assoc zadd_int) |
|
371 |
||
372 |
lemma zless_trans: "[| z1<z2; z2<z3 |] ==> z1 < (z3::int)" |
|
373 |
by (auto simp add: zless_iff_Suc_zadd zadd_assoc zadd_int) |
|
374 |
||
375 |
lemma zless_not_sym: "!!w::int. z<w ==> ~w<z" |
|
376 |
apply (safe dest!: zless_iff_Suc_zadd [THEN iffD1]) |
|
377 |
apply (rule_tac z = z in eq_Abs_Integ, safe) |
|
378 |
apply (simp add: int_def zadd) |
|
379 |
done |
|
380 |
||
381 |
(* [| n<m; ~P ==> m<n |] ==> P *) |
|
382 |
lemmas zless_asym = zless_not_sym [THEN swap, standard] |
|
383 |
||
384 |
lemma zless_not_refl: "!!z::int. ~ z<z" |
|
385 |
apply (rule zless_asym [THEN notI]) |
|
386 |
apply (assumption+) |
|
387 |
done |
|
388 |
||
389 |
(* z<z ==> R *) |
|
390 |
lemmas zless_irrefl = zless_not_refl [THEN notE, standard, elim!] |
|
391 |
||
392 |
||
393 |
(*"Less than" is a linear ordering*) |
|
394 |
lemma zless_linear: |
|
395 |
"z<w | z=w | w<(z::int)" |
|
396 |
apply (unfold zless_def neg_def zdiff_def) |
|
397 |
apply (rule_tac z = z in eq_Abs_Integ) |
|
398 |
apply (rule_tac z = w in eq_Abs_Integ, safe) |
|
399 |
apply (simp add: zadd zminus Image_iff Bex_def) |
|
400 |
apply (rule_tac m1 = "x+ya" and n1 = "xa+y" in less_linear [THEN disjE]) |
|
401 |
apply (force simp add: add_ac)+ |
|
402 |
done |
|
403 |
||
404 |
lemma int_neq_iff: "!!w::int. (w ~= z) = (w<z | z<w)" |
|
405 |
by (cut_tac zless_linear, blast) |
|
406 |
||
407 |
(*** eliminates ~= in premises ***) |
|
408 |
lemmas int_neqE = int_neq_iff [THEN iffD1, THEN disjE, standard] |
|
409 |
||
410 |
lemma int_eq_0_conv [simp]: "(int n = 0) = (n = 0)" |
|
411 |
by (simp add: Zero_int_def) |
|
412 |
||
413 |
lemma zless_int [simp]: "(int m < int n) = (m<n)" |
|
414 |
by (simp add: less_iff_Suc_add zless_iff_Suc_zadd zadd_int) |
|
415 |
||
416 |
lemma int_less_0_conv [simp]: "~ (int k < 0)" |
|
417 |
by (simp add: Zero_int_def) |
|
418 |
||
419 |
lemma zero_less_int_conv [simp]: "(0 < int n) = (0 < n)" |
|
420 |
by (simp add: Zero_int_def) |
|
421 |
||
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
422 |
lemma int_0_less_1: "0 < (1::int)" |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
423 |
by (simp only: Zero_int_def One_int_def One_nat_def zless_int) |
14259 | 424 |
|
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
425 |
lemma int_0_neq_1 [simp]: "0 \<noteq> (1::int)" |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
426 |
by (simp only: Zero_int_def One_int_def One_nat_def int_int_eq) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
427 |
|
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
428 |
|
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
429 |
subsection{*Properties of the @{text "\<le>"} Relation*} |
14259 | 430 |
|
431 |
lemma zle_int [simp]: "(int m <= int n) = (m<=n)" |
|
432 |
by (simp add: zle_def le_def) |
|
433 |
||
434 |
lemma zero_zle_int [simp]: "(0 <= int n)" |
|
435 |
by (simp add: Zero_int_def) |
|
436 |
||
437 |
lemma int_le_0_conv [simp]: "(int n <= 0) = (n = 0)" |
|
438 |
by (simp add: Zero_int_def) |
|
439 |
||
440 |
lemma zle_imp_zless_or_eq: "z <= w ==> z < w | z=(w::int)" |
|
441 |
apply (unfold zle_def) |
|
442 |
apply (cut_tac zless_linear) |
|
443 |
apply (blast elim: zless_asym) |
|
444 |
done |
|
445 |
||
446 |
lemma zless_or_eq_imp_zle: "z<w | z=w ==> z <= (w::int)" |
|
447 |
apply (unfold zle_def) |
|
448 |
apply (cut_tac zless_linear) |
|
449 |
apply (blast elim: zless_asym) |
|
450 |
done |
|
451 |
||
452 |
lemma int_le_less: "(x <= (y::int)) = (x < y | x=y)" |
|
453 |
apply (rule iffI) |
|
454 |
apply (erule zle_imp_zless_or_eq) |
|
455 |
apply (erule zless_or_eq_imp_zle) |
|
456 |
done |
|
457 |
||
458 |
lemma zle_refl: "w <= (w::int)" |
|
459 |
by (simp add: int_le_less) |
|
460 |
||
461 |
(* Axiom 'linorder_linear' of class 'linorder': *) |
|
462 |
lemma zle_linear: "(z::int) <= w | w <= z" |
|
463 |
apply (simp add: int_le_less) |
|
464 |
apply (cut_tac zless_linear, blast) |
|
465 |
done |
|
466 |
||
467 |
(* Axiom 'order_trans of class 'order': *) |
|
468 |
lemma zle_trans: "[| i <= j; j <= k |] ==> i <= (k::int)" |
|
469 |
apply (drule zle_imp_zless_or_eq) |
|
470 |
apply (drule zle_imp_zless_or_eq) |
|
471 |
apply (rule zless_or_eq_imp_zle) |
|
472 |
apply (blast intro: zless_trans) |
|
473 |
done |
|
474 |
||
475 |
lemma zle_anti_sym: "[| z <= w; w <= z |] ==> z = (w::int)" |
|
476 |
apply (drule zle_imp_zless_or_eq) |
|
477 |
apply (drule zle_imp_zless_or_eq) |
|
478 |
apply (blast elim: zless_asym) |
|
479 |
done |
|
480 |
||
481 |
(* Axiom 'order_less_le' of class 'order': *) |
|
482 |
lemma int_less_le: "((w::int) < z) = (w <= z & w ~= z)" |
|
483 |
apply (simp add: zle_def int_neq_iff) |
|
484 |
apply (blast elim!: zless_asym) |
|
485 |
done |
|
486 |
||
14271 | 487 |
instance int :: order |
488 |
proof qed (assumption | rule zle_refl zle_trans zle_anti_sym int_less_le)+ |
|
14259 | 489 |
|
14271 | 490 |
instance int :: plus_ac0 |
491 |
proof qed (rule zadd_commute zadd_assoc zadd_0)+ |
|
14259 | 492 |
|
14271 | 493 |
instance int :: linorder |
494 |
proof qed (rule zle_linear) |
|
495 |
||
14259 | 496 |
|
14341
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
497 |
lemma zadd_left_cancel [simp]: "!!w::int. (z + w' = z + w) = (w' = w)" |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
498 |
by (rule add_left_cancel) |
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
499 |
|
a09441bd4f1e
Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents:
14271
diff
changeset
|
500 |
|
14259 | 501 |
ML |
502 |
{* |
|
503 |
val int_def = thm "int_def"; |
|
504 |
val neg_def = thm "neg_def"; |
|
505 |
val iszero_def = thm "iszero_def"; |
|
506 |
val Zero_int_def = thm "Zero_int_def"; |
|
507 |
val One_int_def = thm "One_int_def"; |
|
508 |
val zadd_def = thm "zadd_def"; |
|
509 |
val zdiff_def = thm "zdiff_def"; |
|
510 |
val zless_def = thm "zless_def"; |
|
511 |
val zle_def = thm "zle_def"; |
|
512 |
val zmult_def = thm "zmult_def"; |
|
513 |
||
514 |
val intrel_iff = thm "intrel_iff"; |
|
515 |
val equiv_intrel = thm "equiv_intrel"; |
|
516 |
val equiv_intrel_iff = thm "equiv_intrel_iff"; |
|
517 |
val intrel_in_integ = thm "intrel_in_integ"; |
|
518 |
val inj_on_Abs_Integ = thm "inj_on_Abs_Integ"; |
|
519 |
val inj_Rep_Integ = thm "inj_Rep_Integ"; |
|
520 |
val inj_int = thm "inj_int"; |
|
521 |
val zminus_congruent = thm "zminus_congruent"; |
|
522 |
val zminus = thm "zminus"; |
|
523 |
val eq_Abs_Integ = thm "eq_Abs_Integ"; |
|
524 |
val zminus_zminus = thm "zminus_zminus"; |
|
525 |
val inj_zminus = thm "inj_zminus"; |
|
526 |
val zminus_0 = thm "zminus_0"; |
|
527 |
val not_neg_int = thm "not_neg_int"; |
|
528 |
val neg_zminus_int = thm "neg_zminus_int"; |
|
529 |
val zadd = thm "zadd"; |
|
530 |
val zminus_zadd_distrib = thm "zminus_zadd_distrib"; |
|
531 |
val zadd_commute = thm "zadd_commute"; |
|
532 |
val zadd_assoc = thm "zadd_assoc"; |
|
533 |
val zadd_left_commute = thm "zadd_left_commute"; |
|
534 |
val zadd_ac = thms "zadd_ac"; |
|
14271 | 535 |
val zmult_ac = thms "zmult_ac"; |
14259 | 536 |
val zadd_int = thm "zadd_int"; |
537 |
val zadd_int_left = thm "zadd_int_left"; |
|
538 |
val int_Suc = thm "int_Suc"; |
|
539 |
val zadd_0 = thm "zadd_0"; |
|
540 |
val zadd_0_right = thm "zadd_0_right"; |
|
541 |
val zadd_zminus_inverse = thm "zadd_zminus_inverse"; |
|
542 |
val zadd_zminus_inverse2 = thm "zadd_zminus_inverse2"; |
|
543 |
val zadd_zminus_cancel = thm "zadd_zminus_cancel"; |
|
544 |
val zminus_zadd_cancel = thm "zminus_zadd_cancel"; |
|
545 |
val zdiff0 = thm "zdiff0"; |
|
546 |
val zdiff0_right = thm "zdiff0_right"; |
|
547 |
val zdiff_self = thm "zdiff_self"; |
|
548 |
val zadd_assoc_cong = thm "zadd_assoc_cong"; |
|
549 |
val zadd_assoc_swap = thm "zadd_assoc_swap"; |
|
550 |
val zmult_congruent2 = thm "zmult_congruent2"; |
|
551 |
val zmult = thm "zmult"; |
|
552 |
val zmult_zminus = thm "zmult_zminus"; |
|
553 |
val zmult_commute = thm "zmult_commute"; |
|
554 |
val zmult_assoc = thm "zmult_assoc"; |
|
555 |
val zadd_zmult_distrib = thm "zadd_zmult_distrib"; |
|
556 |
val zmult_zminus_right = thm "zmult_zminus_right"; |
|
557 |
val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2"; |
|
558 |
val zdiff_zmult_distrib = thm "zdiff_zmult_distrib"; |
|
559 |
val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2"; |
|
560 |
val int_distrib = thms "int_distrib"; |
|
561 |
val zmult_int = thm "zmult_int"; |
|
562 |
val zmult_0 = thm "zmult_0"; |
|
563 |
val zmult_1 = thm "zmult_1"; |
|
564 |
val zmult_0_right = thm "zmult_0_right"; |
|
565 |
val zmult_1_right = thm "zmult_1_right"; |
|
566 |
val zless_iff_Suc_zadd = thm "zless_iff_Suc_zadd"; |
|
567 |
val zless_zadd_Suc = thm "zless_zadd_Suc"; |
|
568 |
val zless_trans = thm "zless_trans"; |
|
569 |
val zless_not_sym = thm "zless_not_sym"; |
|
570 |
val zless_asym = thm "zless_asym"; |
|
571 |
val zless_not_refl = thm "zless_not_refl"; |
|
572 |
val zless_irrefl = thm "zless_irrefl"; |
|
573 |
val zless_linear = thm "zless_linear"; |
|
574 |
val int_neq_iff = thm "int_neq_iff"; |
|
575 |
val int_neqE = thm "int_neqE"; |
|
576 |
val int_int_eq = thm "int_int_eq"; |
|
577 |
val int_eq_0_conv = thm "int_eq_0_conv"; |
|
578 |
val zless_int = thm "zless_int"; |
|
579 |
val int_less_0_conv = thm "int_less_0_conv"; |
|
580 |
val zero_less_int_conv = thm "zero_less_int_conv"; |
|
581 |
val zle_int = thm "zle_int"; |
|
582 |
val zero_zle_int = thm "zero_zle_int"; |
|
583 |
val int_le_0_conv = thm "int_le_0_conv"; |
|
584 |
val zle_imp_zless_or_eq = thm "zle_imp_zless_or_eq"; |
|
585 |
val zless_or_eq_imp_zle = thm "zless_or_eq_imp_zle"; |
|
586 |
val int_le_less = thm "int_le_less"; |
|
587 |
val zle_refl = thm "zle_refl"; |
|
588 |
val zle_linear = thm "zle_linear"; |
|
589 |
val zle_trans = thm "zle_trans"; |
|
590 |
val zle_anti_sym = thm "zle_anti_sym"; |
|
591 |
val int_less_le = thm "int_less_le"; |
|
592 |
val zadd_left_cancel = thm "zadd_left_cancel"; |
|
593 |
*} |
|
594 |
||
5508 | 595 |
end |