81 apply (rule Abs_Integ_cases [of z]) |
79 apply (rule Abs_Integ_cases [of z]) |
82 apply (auto simp add: Integ_def quotient_def) |
80 apply (auto simp add: Integ_def quotient_def) |
83 done |
81 done |
84 |
82 |
85 |
83 |
86 subsubsection{*Integer Unary Negation*} |
84 subsection{*Arithmetic Operations*} |
87 |
85 |
88 lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})" |
86 lemma minus: "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})" |
89 proof - |
87 proof - |
90 have "(\<lambda>(x,y). intrel``{(y,x)}) respects intrel" |
88 have "(\<lambda>(x,y). intrel``{(y,x)}) respects intrel" |
91 by (simp add: congruent_def) |
89 by (simp add: congruent_def) |
92 thus ?thesis |
90 thus ?thesis |
93 by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel]) |
91 by (simp add: minus_int_def UN_equiv_class [OF equiv_intrel]) |
94 qed |
92 qed |
95 |
|
96 lemma zminus_zminus: "- (- z) = (z::int)" |
|
97 by (cases z) (simp add: minus) |
|
98 |
|
99 lemma zminus_0: "- 0 = (0::int)" |
|
100 by (simp add: Zero_int_def minus) |
|
101 |
|
102 |
|
103 subsection{*Integer Addition*} |
|
104 |
93 |
105 lemma add: |
94 lemma add: |
106 "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) = |
95 "Abs_Integ (intrel``{(x,y)}) + Abs_Integ (intrel``{(u,v)}) = |
107 Abs_Integ (intrel``{(x+u, y+v)})" |
96 Abs_Integ (intrel``{(x+u, y+v)})" |
108 proof - |
97 proof - |
111 by (simp add: congruent2_def) |
100 by (simp add: congruent2_def) |
112 thus ?thesis |
101 thus ?thesis |
113 by (simp add: add_int_def UN_UN_split_split_eq |
102 by (simp add: add_int_def UN_UN_split_split_eq |
114 UN_equiv_class2 [OF equiv_intrel equiv_intrel]) |
103 UN_equiv_class2 [OF equiv_intrel equiv_intrel]) |
115 qed |
104 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" |
|
139 apply (simp add: Zero_int_def) |
|
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)" |
|
147 by (cases z, simp add: Zero_int_def minus add) |
|
148 |
|
149 |
|
150 subsection{*Integer Multiplication*} |
|
151 |
105 |
152 text{*Congruence property for multiplication*} |
106 text{*Congruence property for multiplication*} |
153 lemma mult_congruent2: |
107 lemma mult_congruent2: |
154 "(%p1 p2. (%(x,y). (%(u,v). intrel``{(x*u + y*v, x*v + y*u)}) p2) p1) |
108 "(%p1 p2. (%(x,y). (%(u,v). intrel``{(x*u + y*v, x*v + y*u)}) p2) p1) |
155 respects2 intrel" |
109 respects2 intrel" |
160 apply (subgoal_tac "u*y + x*y = w*y + v*y & u*z + x*z = w*z + v*z") |
114 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) |
115 apply (simp add: mult_ac) |
162 apply (simp add: add_mult_distrib [symmetric]) |
116 apply (simp add: add_mult_distrib [symmetric]) |
163 done |
117 done |
164 |
118 |
165 |
|
166 lemma mult: |
119 lemma mult: |
167 "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) = |
120 "Abs_Integ((intrel``{(x,y)})) * Abs_Integ((intrel``{(u,v)})) = |
168 Abs_Integ(intrel `` {(x*u + y*v, x*v + y*u)})" |
121 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 |
122 by (simp add: mult_int_def UN_UN_split_split_eq mult_congruent2 |
170 UN_equiv_class2 [OF equiv_intrel equiv_intrel]) |
123 UN_equiv_class2 [OF equiv_intrel equiv_intrel]) |
171 |
124 |
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" |
|
199 by (cases z, simp add: One_int_def mult) |
|
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}*} |
125 text{*The integers form a @{text comm_ring_1}*} |
206 instance int :: comm_ring_1 |
126 instance int :: comm_ring_1 |
207 proof |
127 proof |
208 fix i j k :: int |
128 fix i j k :: int |
209 show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc) |
129 show "(i + j) + k = i + (j + k)" |
210 show "i + j = j + i" by (simp add: zadd_commute) |
130 by (cases i, cases j, cases k) (simp add: add add_assoc) |
211 show "0 + i = i" by (rule zadd_0) |
131 show "i + j = j + i" |
212 show "- i + i = 0" by (rule zadd_zminus_inverse2) |
132 by (cases i, cases j) (simp add: add_ac add) |
213 show "i - j = i + (-j)" by (simp add: diff_int_def) |
133 show "0 + i = i" |
214 show "(i * j) * k = i * (j * k)" by (rule zmult_assoc) |
134 by (cases i) (simp add: Zero_int_def add) |
215 show "i * j = j * i" by (rule zmult_commute) |
135 show "- i + i = 0" |
216 show "1 * i = i" by (rule zmult_1) |
136 by (cases i) (simp add: Zero_int_def minus add) |
217 show "(i + j) * k = i * k + j * k" by (simp add: int_distrib) |
137 show "i - j = i + - j" |
218 show "0 \<noteq> (1::int)" by (simp add: Zero_int_def One_int_def) |
138 by (simp add: diff_int_def) |
|
139 show "(i * j) * k = i * (j * k)" |
|
140 by (cases i, cases j, cases k) (simp add: mult ring_eq_simps) |
|
141 show "i * j = j * i" |
|
142 by (cases i, cases j) (simp add: mult ring_eq_simps) |
|
143 show "1 * i = i" |
|
144 by (cases i) (simp add: One_int_def mult) |
|
145 show "(i + j) * k = i * k + j * k" |
|
146 by (cases i, cases j, cases k) (simp add: add mult ring_eq_simps) |
|
147 show "0 \<noteq> (1::int)" |
|
148 by (simp add: Zero_int_def One_int_def) |
219 qed |
149 qed |
220 |
150 |
221 abbreviation |
151 abbreviation |
222 int :: "nat \<Rightarrow> int" |
152 int :: "nat \<Rightarrow> int" |
223 where |
153 where |
235 |
165 |
236 lemma less: |
166 lemma less: |
237 "(Abs_Integ(intrel``{(x,y)}) < Abs_Integ(intrel``{(u,v)})) = (x+v < u+y)" |
167 "(Abs_Integ(intrel``{(x,y)}) < Abs_Integ(intrel``{(u,v)})) = (x+v < u+y)" |
238 by (simp add: less_int_def le order_less_le) |
168 by (simp add: less_int_def le order_less_le) |
239 |
169 |
240 lemma zle_refl: "w \<le> (w::int)" |
|
241 by (cases w, simp add: le) |
|
242 |
|
243 lemma zle_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::int)" |
|
244 by (cases i, cases j, cases k, simp add: le) |
|
245 |
|
246 lemma zle_anti_sym: "[| z \<le> w; w \<le> z |] ==> z = (w::int)" |
|
247 by (cases w, cases z, simp add: le) |
|
248 |
|
249 instance int :: order |
|
250 by intro_classes |
|
251 (assumption | |
|
252 rule zle_refl zle_trans zle_anti_sym less_int_def [THEN meta_eq_to_obj_eq])+ |
|
253 |
|
254 lemma zle_linear: "(z::int) \<le> w \<or> w \<le> z" |
|
255 by (cases z, cases w) (simp add: le linorder_linear) |
|
256 |
|
257 instance int :: linorder |
170 instance int :: linorder |
258 by intro_classes (rule zle_linear) |
171 proof |
259 |
172 fix i j k :: int |
260 lemmas zless_linear = linorder_less_linear [where 'a = int] |
173 show "(i < j) = (i \<le> j \<and> i \<noteq> j)" |
261 |
174 by (simp add: less_int_def) |
262 |
175 show "i \<le> i" |
263 lemma int_0_less_1: "0 < (1::int)" |
176 by (cases i) (simp add: le) |
264 by (simp add: Zero_int_def One_int_def less) |
177 show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> i \<le> k" |
265 |
178 by (cases i, cases j, cases k) (simp add: le) |
266 lemma int_0_neq_1 [simp]: "0 \<noteq> (1::int)" |
179 show "i \<le> j \<Longrightarrow> j \<le> i \<Longrightarrow> i = j" |
267 by (rule int_0_less_1 [THEN less_imp_neq]) |
180 by (cases i, cases j) (simp add: le) |
268 |
181 show "i \<le> j \<or> j \<le> i" |
269 |
182 by (cases i, cases j) (simp add: le linorder_linear) |
270 subsection{*Monotonicity results*} |
183 qed |
271 |
184 |
272 instance int :: pordered_cancel_ab_semigroup_add |
185 instance int :: pordered_cancel_ab_semigroup_add |
273 proof |
186 proof |
274 fix a b c :: int |
187 fix i j k :: int |
275 show "a \<le> b \<Longrightarrow> c + a \<le> c + b" |
188 show "i \<le> j \<Longrightarrow> k + i \<le> k + j" |
276 by (cases a, cases b, cases c, simp add: le add) |
189 by (cases i, cases j, cases k) (simp add: le add) |
277 qed |
190 qed |
278 |
191 |
279 lemma zadd_left_mono: "i \<le> j ==> k + i \<le> k + (j::int)" |
192 text{*Strict Monotonicity of Multiplication*} |
280 by (rule add_left_mono) |
|
281 |
|
282 lemma zadd_strict_right_mono: "i < j ==> i + k < j + (k::int)" |
|
283 by (rule add_strict_right_mono) |
|
284 |
|
285 lemma zadd_zless_mono: "[| w'<w; z'\<le>z |] ==> w' + z' < w + (z::int)" |
|
286 by (rule add_less_le_mono) |
|
287 |
|
288 |
|
289 subsection{*Strict Monotonicity of Multiplication*} |
|
290 |
193 |
291 text{*strict, in 1st argument; proof is by induction on k>0*} |
194 text{*strict, in 1st argument; proof is by induction on k>0*} |
292 lemma zmult_zless_mono2_lemma: |
195 lemma zmult_zless_mono2_lemma: |
293 "(i::int)<j ==> 0<k ==> of_nat k * i < of_nat k * j" |
196 "(i::int)<j ==> 0<k ==> int k * i < int k * j" |
294 apply (induct "k", simp) |
197 apply (induct "k", simp) |
295 apply (simp add: left_distrib) |
198 apply (simp add: left_distrib) |
296 apply (case_tac "k=0") |
199 apply (case_tac "k=0") |
297 apply (simp_all add: add_strict_mono) |
200 apply (simp_all add: add_strict_mono) |
298 done |
201 done |
299 |
202 |
300 lemma zero_le_imp_eq_int: "(0::int) \<le> k ==> \<exists>n. k = of_nat n" |
203 lemma zero_le_imp_eq_int: "(0::int) \<le> k ==> \<exists>n. k = int n" |
301 apply (cases k) |
204 apply (cases k) |
302 apply (auto simp add: le add int_def Zero_int_def) |
205 apply (auto simp add: le add int_def Zero_int_def) |
303 apply (rule_tac x="x-y" in exI, simp) |
206 apply (rule_tac x="x-y" in exI, simp) |
304 done |
207 done |
305 |
208 |
306 lemma zero_less_imp_eq_int: "(0::int) < k ==> \<exists>n>0. k = of_nat n" |
209 lemma zero_less_imp_eq_int: "(0::int) < k ==> \<exists>n>0. k = int n" |
307 apply (cases k) |
210 apply (cases k) |
308 apply (simp add: less int_def Zero_int_def) |
211 apply (simp add: less int_def Zero_int_def) |
309 apply (rule_tac x="x-y" in exI, simp) |
212 apply (rule_tac x="x-y" in exI, simp) |
310 done |
213 done |
311 |
214 |
325 |
228 |
326 text{*The integers form an ordered integral domain*} |
229 text{*The integers form an ordered integral domain*} |
327 instance int :: ordered_idom |
230 instance int :: ordered_idom |
328 proof |
231 proof |
329 fix i j k :: int |
232 fix i j k :: int |
330 show "i < j ==> 0 < k ==> k * i < k * j" by (rule zmult_zless_mono2) |
233 show "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j" |
331 show "\<bar>i\<bar> = (if i < 0 then -i else i)" by (simp only: zabs_def) |
234 by (rule zmult_zless_mono2) |
|
235 show "\<bar>i\<bar> = (if i < 0 then -i else i)" |
|
236 by (simp only: zabs_def) |
332 qed |
237 qed |
333 |
238 |
334 lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z" |
239 lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z" |
335 apply (cases w, cases z) |
240 apply (cases w, cases z) |
336 apply (simp add: less le add One_int_def) |
241 apply (simp add: less le add One_int_def) |
445 |
350 |
446 lemma negative_eq_positive [simp]: "(- int n = int m) = (n = 0 & m = 0)" |
351 lemma negative_eq_positive [simp]: "(- int n = int m) = (n = 0 & m = 0)" |
447 by (force simp add: order_eq_iff [of "- int n"] int_zle_neg) |
352 by (force simp add: order_eq_iff [of "- int n"] int_zle_neg) |
448 |
353 |
449 lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)" |
354 lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)" |
450 proof (cases w, cases z, simp add: le add int_def) |
355 proof - |
451 fix a b c d |
356 have "(w \<le> z) = (0 \<le> z - w)" |
452 assume "w = Abs_Integ (intrel `` {(a,b)})" "z = Abs_Integ (intrel `` {(c,d)})" |
357 by (simp only: le_diff_eq add_0_left) |
453 show "(a+d \<le> c+b) = (\<exists>n. c+b = a+n+d)" |
358 also have "\<dots> = (\<exists>n. z - w = int n)" |
454 proof |
359 by (auto elim: zero_le_imp_eq_int) |
455 assume "a + d \<le> c + b" |
360 also have "\<dots> = (\<exists>n. z = w + int n)" |
456 thus "\<exists>n. c + b = a + n + d" |
361 by (simp only: group_eq_simps) |
457 by (auto intro!: exI [where x="c+b - (a+d)"]) |
362 finally show ?thesis . |
458 next |
363 qed |
459 assume "\<exists>n. c + b = a + n + d" |
364 |
460 then obtain n where "c + b = a + n + d" .. |
365 lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z" |
461 thus "a + d \<le> c + b" by arith |
366 by simp |
462 qed |
367 |
463 qed |
368 lemma int_Suc: "int (Suc m) = 1 + (int m)" |
|
369 by simp |
|
370 |
|
371 lemma int_Suc0_eq_1: "int (Suc 0) = 1" |
|
372 by simp |
464 |
373 |
465 lemma abs_of_nat [simp]: "\<bar>of_nat n::'a::ordered_idom\<bar> = of_nat n" |
374 lemma abs_of_nat [simp]: "\<bar>of_nat n::'a::ordered_idom\<bar> = of_nat n" |
466 by (rule of_nat_0_le_iff [THEN abs_of_nonneg]) (* belongs in Nat.thy *) |
375 by (rule of_nat_0_le_iff [THEN abs_of_nonneg]) (* belongs in Nat.thy *) |
467 |
376 |
468 text{*This version is proved for all ordered rings, not just integers! |
377 text{*This version is proved for all ordered rings, not just integers! |
783 apply (cases "z < 0", blast dest!: negD) |
695 apply (cases "z < 0", blast dest!: negD) |
784 apply (simp add: linorder_not_less del: of_nat_Suc) |
696 apply (simp add: linorder_not_less del: of_nat_Suc) |
785 apply (blast dest: nat_0_le [THEN sym]) |
697 apply (blast dest: nat_0_le [THEN sym]) |
786 done |
698 done |
787 |
699 |
788 theorem int_induct'[induct type: int, case_names nonneg neg]: |
700 theorem int_induct [induct type: int, case_names nonneg neg]: |
789 "[|!! n. P (int n); !!n. P (- (int (Suc n))) |] ==> P z" |
701 "[|!! n. P (int n); !!n. P (- (int (Suc n))) |] ==> P z" |
790 by (cases z rule: int_cases) auto |
702 by (cases z rule: int_cases) auto |
791 |
703 |
792 text{*Contributed by Brian Huffman*} |
704 text{*Contributed by Brian Huffman*} |
793 theorem int_diff_cases [case_names diff]: |
705 theorem int_diff_cases [case_names diff]: |
795 apply (cases z rule: eq_Abs_Integ) |
707 apply (cases z rule: eq_Abs_Integ) |
796 apply (rule_tac m=x and n=y in prem) |
708 apply (rule_tac m=x and n=y in prem) |
797 apply (simp add: int_def diff_def minus add) |
709 apply (simp add: int_def diff_def minus add) |
798 done |
710 done |
799 |
711 |
800 lemma of_nat_nat: "0 \<le> z ==> of_nat (nat z) = of_int z" |
|
801 by (cases z rule: eq_Abs_Integ, simp add: nat le of_int Zero_int_def) |
|
802 |
|
803 |
712 |
804 subsection {* Legacy theorems *} |
713 subsection {* Legacy theorems *} |
805 |
714 |
806 lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z" |
715 lemmas zminus_zminus = minus_minus [where 'a=int] |
807 by simp |
716 lemmas zminus_0 = minus_zero [where 'a=int] |
808 |
717 lemmas zminus_zadd_distrib = minus_add_distrib [where 'a=int] |
809 lemma int_Suc: "int (Suc m) = 1 + (int m)" |
718 lemmas zadd_commute = add_commute [where 'a=int] |
810 by simp |
719 lemmas zadd_assoc = add_assoc [where 'a=int] |
811 |
720 lemmas zadd_left_commute = add_left_commute [where 'a=int] |
812 lemma int_Suc0_eq_1: "int (Suc 0) = 1" |
721 lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute |
813 by simp |
722 lemmas zmult_ac = OrderedGroup.mult_ac |
|
723 lemmas zadd_0 = OrderedGroup.add_0_left [where 'a=int] |
|
724 lemmas zadd_0_right = OrderedGroup.add_0_left [where 'a=int] |
|
725 lemmas zadd_zminus_inverse2 = left_minus [where 'a=int] |
|
726 lemmas zmult_zminus = mult_minus_left [where 'a=int] |
|
727 lemmas zmult_commute = mult_commute [where 'a=int] |
|
728 lemmas zmult_assoc = mult_assoc [where 'a=int] |
|
729 lemmas zadd_zmult_distrib = left_distrib [where 'a=int] |
|
730 lemmas zadd_zmult_distrib2 = right_distrib [where 'a=int] |
|
731 lemmas zdiff_zmult_distrib = left_diff_distrib [where 'a=int] |
|
732 lemmas zdiff_zmult_distrib2 = right_diff_distrib [where 'a=int] |
|
733 |
|
734 lemmas int_distrib = |
|
735 zadd_zmult_distrib zadd_zmult_distrib2 |
|
736 zdiff_zmult_distrib zdiff_zmult_distrib2 |
|
737 |
|
738 lemmas zmult_1 = mult_1_left [where 'a=int] |
|
739 lemmas zmult_1_right = mult_1_right [where 'a=int] |
|
740 |
|
741 lemmas zle_refl = order_refl [where 'a=int] |
|
742 lemmas zle_trans = order_trans [where 'a=int] |
|
743 lemmas zle_anti_sym = order_antisym [where 'a=int] |
|
744 lemmas zle_linear = linorder_linear [where 'a=int] |
|
745 lemmas zless_linear = linorder_less_linear [where 'a = int] |
|
746 |
|
747 lemmas zadd_left_mono = add_left_mono [where 'a=int] |
|
748 lemmas zadd_strict_right_mono = add_strict_right_mono [where 'a=int] |
|
749 lemmas zadd_zless_mono = add_less_le_mono [where 'a=int] |
|
750 |
|
751 lemmas int_0_less_1 = zero_less_one [where 'a=int] |
|
752 lemmas int_0_neq_1 = zero_neq_one [where 'a=int] |
814 |
753 |
815 lemmas inj_int = inj_of_nat [where 'a=int] |
754 lemmas inj_int = inj_of_nat [where 'a=int] |
816 lemmas int_int_eq = of_nat_eq_iff [where 'a=int] |
755 lemmas int_int_eq = of_nat_eq_iff [where 'a=int] |
817 lemmas zadd_int = of_nat_add [where 'a=int, symmetric] |
756 lemmas zadd_int = of_nat_add [where 'a=int, symmetric] |
818 lemmas int_mult = of_nat_mult [where 'a=int] |
757 lemmas int_mult = of_nat_mult [where 'a=int] |