| author | paulson |
| Tue, 24 Feb 1998 11:35:33 +0100 | |
| changeset 4648 | f04da668581c |
| parent 4152 | 451104c223e2 |
| child 5068 | fb28eaa07e01 |
| permissions | -rw-r--r-- |
| 1461 | 1 |
(* Title: ZF/ex/integ.ML |
| 0 | 2 |
ID: $Id$ |
| 1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
| 0 | 4 |
Copyright 1993 University of Cambridge |
5 |
||
| 1614 | 6 |
The integers as equivalence classes over nat*nat. |
| 0 | 7 |
|
8 |
Could also prove... |
|
9 |
"znegative(z) ==> $# zmagnitude(z) = $~ z" |
|
10 |
"~ znegative(z) ==> $# zmagnitude(z) = z" |
|
11 |
$< is a linear ordering |
|
12 |
$+ and $* are monotonic wrt $< |
|
13 |
*) |
|
14 |
||
|
7
268f93ab3bc4
Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents:
0
diff
changeset
|
15 |
open Integ; |
| 0 | 16 |
|
17 |
(*** Proving that intrel is an equivalence relation ***) |
|
18 |
||
19 |
(*By luck, requires no typing premises for y1, y2,y3*) |
|
20 |
val eqa::eqb::prems = goal Arith.thy |
|
21 |
"[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2; \ |
|
22 |
\ x1: nat; x2: nat; x3: nat |] ==> x1 #+ y3 = x3 #+ y1"; |
|
23 |
by (res_inst_tac [("k","x2")] add_left_cancel 1);
|
|
| 438 | 24 |
by (resolve_tac prems 2); |
25 |
by (rtac (add_left_commute RS trans) 1 THEN typechk_tac prems); |
|
| 2034 | 26 |
by (stac eqb 1); |
| 438 | 27 |
by (rtac (add_left_commute RS trans) 1 THEN typechk_tac prems); |
| 2034 | 28 |
by (stac eqa 1); |
| 438 | 29 |
by (rtac (add_left_commute) 1 THEN typechk_tac prems); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
30 |
qed "integ_trans_lemma"; |
| 0 | 31 |
|
32 |
(** Natural deduction for intrel **) |
|
33 |
||
| 438 | 34 |
goalw Integ.thy [intrel_def] |
35 |
"<<x1,y1>,<x2,y2>>: intrel <-> \ |
|
36 |
\ x1: nat & y1: nat & x2: nat & y2: nat & x1#+y2 = x2#+y1"; |
|
| 2469 | 37 |
by (Fast_tac 1); |
| 760 | 38 |
qed "intrel_iff"; |
| 438 | 39 |
|
40 |
goalw Integ.thy [intrel_def] |
|
41 |
"!!x1 x2. [| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ |
|
42 |
\ <<x1,y1>,<x2,y2>>: intrel"; |
|
| 4091 | 43 |
by (fast_tac (claset() addIs prems) 1); |
| 760 | 44 |
qed "intrelI"; |
| 0 | 45 |
|
46 |
(*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*) |
|
47 |
goalw Integ.thy [intrel_def] |
|
48 |
"p: intrel --> (EX x1 y1 x2 y2. \ |
|
49 |
\ p = <<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1 & \ |
|
50 |
\ x1: nat & y1: nat & x2: nat & y2: nat)"; |
|
| 2469 | 51 |
by (Fast_tac 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
52 |
qed "intrelE_lemma"; |
| 0 | 53 |
|
54 |
val [major,minor] = goal Integ.thy |
|
55 |
"[| p: intrel; \ |
|
56 |
\ !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>; x1#+y2 = x2#+y1; \ |
|
57 |
\ x1: nat; y1: nat; x2: nat; y2: nat |] ==> Q |] \ |
|
58 |
\ ==> Q"; |
|
59 |
by (cut_facts_tac [major RS (intrelE_lemma RS mp)] 1); |
|
60 |
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1)); |
|
| 760 | 61 |
qed "intrelE"; |
| 0 | 62 |
|
| 2469 | 63 |
AddSIs [intrelI]; |
64 |
AddSEs [intrelE]; |
|
| 0 | 65 |
|
| 438 | 66 |
goalw Integ.thy [equiv_def, refl_def, sym_def, trans_def] |
67 |
"equiv(nat*nat, intrel)"; |
|
| 4091 | 68 |
by (fast_tac (claset() addSEs [sym, integ_trans_lemma]) 1); |
| 760 | 69 |
qed "equiv_intrel"; |
| 0 | 70 |
|
71 |
||
| 2469 | 72 |
Addsimps [equiv_intrel RS eq_equiv_class_iff, intrel_iff, |
73 |
add_0_right, add_succ_right]; |
|
74 |
Addcongs [conj_cong]; |
|
| 0 | 75 |
|
| 438 | 76 |
val eq_intrelD = equiv_intrel RSN (2,eq_equiv_class); |
| 0 | 77 |
|
78 |
(** znat: the injection from nat to integ **) |
|
79 |
||
| 438 | 80 |
goalw Integ.thy [integ_def,quotient_def,znat_def] |
81 |
"!!m. m : nat ==> $#m : integ"; |
|
| 4091 | 82 |
by (fast_tac (claset() addSIs [nat_0I]) 1); |
| 760 | 83 |
qed "znat_type"; |
| 0 | 84 |
|
| 438 | 85 |
goalw Integ.thy [znat_def] |
86 |
"!!m n. [| $#m = $#n; n: nat |] ==> m=n"; |
|
87 |
by (dtac eq_intrelD 1); |
|
88 |
by (typechk_tac [nat_0I, SigmaI]); |
|
| 2469 | 89 |
by (Asm_full_simp_tac 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
90 |
qed "znat_inject"; |
| 0 | 91 |
|
92 |
||
93 |
(**** zminus: unary negation on integ ****) |
|
94 |
||
95 |
goalw Integ.thy [congruent_def] |
|
| 1614 | 96 |
"congruent(intrel, %<x,y>. intrel``{<y,x>})";
|
| 4152 | 97 |
by Safe_tac; |
| 4091 | 98 |
by (asm_full_simp_tac (simpset() addsimps add_ac) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
99 |
qed "zminus_congruent"; |
| 0 | 100 |
|
101 |
(*Resolve th against the corresponding facts for zminus*) |
|
102 |
val zminus_ize = RSLIST [equiv_intrel, zminus_congruent]; |
|
103 |
||
| 438 | 104 |
goalw Integ.thy [integ_def,zminus_def] |
105 |
"!!z. z : integ ==> $~z : integ"; |
|
106 |
by (typechk_tac [split_type, SigmaI, zminus_ize UN_equiv_class_type, |
|
| 1461 | 107 |
quotientI]); |
| 760 | 108 |
qed "zminus_type"; |
| 0 | 109 |
|
| 438 | 110 |
goalw Integ.thy [integ_def,zminus_def] |
111 |
"!!z w. [| $~z = $~w; z: integ; w: integ |] ==> z=w"; |
|
112 |
by (etac (zminus_ize UN_equiv_class_inject) 1); |
|
| 4152 | 113 |
by Safe_tac; |
| 438 | 114 |
(*The setloop is only needed because assumptions are in the wrong order!*) |
| 4091 | 115 |
by (asm_full_simp_tac (simpset() addsimps add_ac |
| 1461 | 116 |
setloop dtac eq_intrelD) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
117 |
qed "zminus_inject"; |
| 0 | 118 |
|
| 438 | 119 |
goalw Integ.thy [zminus_def] |
120 |
"!!x y.[| x: nat; y: nat |] ==> $~ (intrel``{<x,y>}) = intrel `` {<y,x>}";
|
|
| 4091 | 121 |
by (asm_simp_tac (simpset() addsimps [zminus_ize UN_equiv_class, SigmaI]) 1); |
| 760 | 122 |
qed "zminus"; |
| 0 | 123 |
|
124 |
goalw Integ.thy [integ_def] "!!z. z : integ ==> $~ ($~ z) = z"; |
|
125 |
by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1)); |
|
| 4091 | 126 |
by (asm_simp_tac (simpset() addsimps [zminus]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
127 |
qed "zminus_zminus"; |
| 0 | 128 |
|
129 |
goalw Integ.thy [integ_def, znat_def] "$~ ($#0) = $#0"; |
|
| 4091 | 130 |
by (simp_tac (simpset() addsimps [zminus]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
131 |
qed "zminus_0"; |
| 0 | 132 |
|
133 |
||
134 |
(**** znegative: the test for negative integers ****) |
|
135 |
||
| 1021 | 136 |
(*No natural number is negative!*) |
| 438 | 137 |
goalw Integ.thy [znegative_def, znat_def] "~ znegative($# n)"; |
| 4152 | 138 |
by Safe_tac; |
| 1021 | 139 |
by (dres_inst_tac [("psi", "?lhs=?rhs")] asm_rl 1);
|
140 |
by (dres_inst_tac [("psi", "?lhs<?rhs")] asm_rl 1);
|
|
| 4091 | 141 |
by (fast_tac (claset() addss |
142 |
(simpset() addsimps [add_le_self2 RS le_imp_not_lt])) 1); |
|
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
143 |
qed "not_znegative_znat"; |
| 0 | 144 |
|
|
29
4ec9b266ccd1
Modification of examples for the new operators, < and le.
lcp
parents:
16
diff
changeset
|
145 |
goalw Integ.thy [znegative_def, znat_def] |
|
4ec9b266ccd1
Modification of examples for the new operators, < and le.
lcp
parents:
16
diff
changeset
|
146 |
"!!n. n: nat ==> znegative($~ $# succ(n))"; |
| 4091 | 147 |
by (asm_simp_tac (simpset() addsimps [zminus]) 1); |
| 0 | 148 |
by (REPEAT |
|
29
4ec9b266ccd1
Modification of examples for the new operators, < and le.
lcp
parents:
16
diff
changeset
|
149 |
(ares_tac [refl, exI, conjI, nat_0_le, |
| 1461 | 150 |
refl RS intrelI RS imageI, consI1, nat_0I, nat_succI] 1)); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
151 |
qed "znegative_zminus_znat"; |
| 0 | 152 |
|
153 |
||
154 |
(**** zmagnitude: magnitide of an integer, as a natural number ****) |
|
155 |
||
156 |
goalw Integ.thy [congruent_def] |
|
| 1614 | 157 |
"congruent(intrel, %<x,y>. (y#-x) #+ (x#-y))"; |
| 4152 | 158 |
by Safe_tac; |
| 2469 | 159 |
by (ALLGOALS Asm_simp_tac); |
| 0 | 160 |
by (etac rev_mp 1); |
| 438 | 161 |
by (res_inst_tac [("m","x1"),("n","y1")] diff_induct 1 THEN
|
162 |
REPEAT (assume_tac 1)); |
|
| 2469 | 163 |
by (Asm_simp_tac 3); |
| 438 | 164 |
by (asm_simp_tac (*this one's very sensitive to order of rewrites*) |
| 4091 | 165 |
(simpset() delsimps [add_succ_right] |
| 2469 | 166 |
addsimps [diff_add_inverse,diff_add_0]) 2); |
167 |
by (Asm_simp_tac 1); |
|
| 0 | 168 |
by (rtac impI 1); |
169 |
by (etac subst 1); |
|
| 438 | 170 |
by (res_inst_tac [("m1","x")] (add_commute RS ssubst) 1 THEN
|
171 |
REPEAT (assume_tac 1)); |
|
| 4091 | 172 |
by (asm_simp_tac (simpset() addsimps [diff_add_inverse, diff_add_0]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
173 |
qed "zmagnitude_congruent"; |
| 0 | 174 |
|
| 438 | 175 |
|
| 0 | 176 |
(*Resolve th against the corresponding facts for zmagnitude*) |
177 |
val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent]; |
|
178 |
||
| 438 | 179 |
goalw Integ.thy [integ_def,zmagnitude_def] |
180 |
"!!z. z : integ ==> zmagnitude(z) : nat"; |
|
181 |
by (typechk_tac [split_type, zmagnitude_ize UN_equiv_class_type, |
|
| 1461 | 182 |
add_type, diff_type]); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
183 |
qed "zmagnitude_type"; |
| 0 | 184 |
|
| 438 | 185 |
goalw Integ.thy [zmagnitude_def] |
186 |
"!!x y. [| x: nat; y: nat |] ==> \ |
|
187 |
\ zmagnitude (intrel``{<x,y>}) = (y #- x) #+ (x #- y)";
|
|
| 2469 | 188 |
by (asm_simp_tac |
| 4091 | 189 |
(simpset() addsimps [zmagnitude_ize UN_equiv_class, SigmaI]) 1); |
| 760 | 190 |
qed "zmagnitude"; |
| 0 | 191 |
|
|
29
4ec9b266ccd1
Modification of examples for the new operators, < and le.
lcp
parents:
16
diff
changeset
|
192 |
goalw Integ.thy [znat_def] |
|
4ec9b266ccd1
Modification of examples for the new operators, < and le.
lcp
parents:
16
diff
changeset
|
193 |
"!!n. n: nat ==> zmagnitude($# n) = n"; |
| 4091 | 194 |
by (asm_simp_tac (simpset() addsimps [zmagnitude]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
195 |
qed "zmagnitude_znat"; |
| 0 | 196 |
|
|
29
4ec9b266ccd1
Modification of examples for the new operators, < and le.
lcp
parents:
16
diff
changeset
|
197 |
goalw Integ.thy [znat_def] |
|
4ec9b266ccd1
Modification of examples for the new operators, < and le.
lcp
parents:
16
diff
changeset
|
198 |
"!!n. n: nat ==> zmagnitude($~ $# n) = n"; |
| 4091 | 199 |
by (asm_simp_tac (simpset() addsimps [zmagnitude, zminus]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
200 |
qed "zmagnitude_zminus_znat"; |
| 0 | 201 |
|
202 |
||
203 |
(**** zadd: addition on integ ****) |
|
204 |
||
205 |
(** Congruence property for addition **) |
|
206 |
||
207 |
goalw Integ.thy [congruent2_def] |
|
| 1461 | 208 |
"congruent2(intrel, %z1 z2. \ |
|
1110
756aa2e81f6e
Changed some definitions and proofs to use pattern-matching.
lcp
parents:
1021
diff
changeset
|
209 |
\ let <x1,y1>=z1; <x2,y2>=z2 \ |
|
756aa2e81f6e
Changed some definitions and proofs to use pattern-matching.
lcp
parents:
1021
diff
changeset
|
210 |
\ in intrel``{<x1#+x2, y1#+y2>})";
|
| 0 | 211 |
(*Proof via congruent2_commuteI seems longer*) |
| 4152 | 212 |
by Safe_tac; |
| 4091 | 213 |
by (asm_simp_tac (simpset() addsimps [add_assoc, Let_def]) 1); |
| 438 | 214 |
(*The rest should be trivial, but rearranging terms is hard; |
215 |
add_ac does not help rewriting with the assumptions.*) |
|
216 |
by (res_inst_tac [("m1","x1a")] (add_left_commute RS ssubst) 1);
|
|
217 |
by (res_inst_tac [("m1","x2a")] (add_left_commute RS ssubst) 3);
|
|
| 0 | 218 |
by (typechk_tac [add_type]); |
| 4091 | 219 |
by (asm_simp_tac (simpset() addsimps [add_assoc RS sym]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
220 |
qed "zadd_congruent2"; |
| 0 | 221 |
|
222 |
(*Resolve th against the corresponding facts for zadd*) |
|
223 |
val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2]; |
|
224 |
||
| 438 | 225 |
goalw Integ.thy [integ_def,zadd_def] |
226 |
"!!z w. [| z: integ; w: integ |] ==> z $+ w : integ"; |
|
|
1110
756aa2e81f6e
Changed some definitions and proofs to use pattern-matching.
lcp
parents:
1021
diff
changeset
|
227 |
by (rtac (zadd_ize UN_equiv_class_type2) 1); |
| 4091 | 228 |
by (simp_tac (simpset() addsimps [Let_def]) 3); |
|
1110
756aa2e81f6e
Changed some definitions and proofs to use pattern-matching.
lcp
parents:
1021
diff
changeset
|
229 |
by (REPEAT (ares_tac [split_type, add_type, quotientI, SigmaI] 1)); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
230 |
qed "zadd_type"; |
| 0 | 231 |
|
| 438 | 232 |
goalw Integ.thy [zadd_def] |
| 1461 | 233 |
"!!x1 y1. [| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ |
234 |
\ (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) = \
|
|
235 |
\ intrel `` {<x1#+x2, y1#+y2>}";
|
|
| 4091 | 236 |
by (asm_simp_tac (simpset() addsimps [zadd_ize UN_equiv_class2, SigmaI]) 1); |
237 |
by (simp_tac (simpset() addsimps [Let_def]) 1); |
|
| 760 | 238 |
qed "zadd"; |
| 0 | 239 |
|
240 |
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $+ z = z"; |
|
| 438 | 241 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 242 |
by (asm_simp_tac (simpset() addsimps [zadd]) 1); |
| 760 | 243 |
qed "zadd_0"; |
| 0 | 244 |
|
245 |
goalw Integ.thy [integ_def] |
|
246 |
"!!z w. [| z: integ; w: integ |] ==> $~ (z $+ w) = $~ z $+ $~ w"; |
|
| 438 | 247 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 248 |
by (asm_simp_tac (simpset() addsimps [zminus,zadd]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
249 |
qed "zminus_zadd_distrib"; |
| 0 | 250 |
|
251 |
goalw Integ.thy [integ_def] |
|
252 |
"!!z w. [| z: integ; w: integ |] ==> z $+ w = w $+ z"; |
|
| 438 | 253 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 254 |
by (asm_simp_tac (simpset() addsimps (add_ac @ [zadd])) 1); |
| 760 | 255 |
qed "zadd_commute"; |
| 0 | 256 |
|
257 |
goalw Integ.thy [integ_def] |
|
258 |
"!!z1 z2 z3. [| z1: integ; z2: integ; z3: integ |] ==> \ |
|
259 |
\ (z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)"; |
|
| 438 | 260 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 0 | 261 |
(*rewriting is much faster without intrel_iff, etc.*) |
| 4091 | 262 |
by (asm_simp_tac (simpset() addsimps [zadd, add_assoc]) 1); |
| 760 | 263 |
qed "zadd_assoc"; |
| 0 | 264 |
|
|
904
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
265 |
(*For AC rewriting*) |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
266 |
qed_goal "zadd_left_commute" Integ.thy |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
267 |
"!!z1 z2 z3. [| z1:integ; z2:integ; z3: integ |] ==> \ |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
268 |
\ z1$+(z2$+z3) = z2$+(z1$+z3)" |
| 4091 | 269 |
(fn _ => [asm_simp_tac (simpset() addsimps [zadd_assoc RS sym, zadd_commute]) 1]); |
|
904
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
270 |
|
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
271 |
(*Integer addition is an AC operator*) |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
272 |
val zadd_ac = [zadd_assoc, zadd_commute, zadd_left_commute]; |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
273 |
|
| 438 | 274 |
goalw Integ.thy [znat_def] |
275 |
"!!m n. [| m: nat; n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)"; |
|
| 4091 | 276 |
by (asm_simp_tac (simpset() addsimps [zadd]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
277 |
qed "znat_add"; |
| 0 | 278 |
|
279 |
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> z $+ ($~ z) = $#0"; |
|
| 438 | 280 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 281 |
by (asm_simp_tac (simpset() addsimps [zminus, zadd, add_commute]) 1); |
| 760 | 282 |
qed "zadd_zminus_inverse"; |
| 0 | 283 |
|
| 438 | 284 |
goal Integ.thy "!!z. z : integ ==> ($~ z) $+ z = $#0"; |
285 |
by (asm_simp_tac |
|
| 4091 | 286 |
(simpset() addsimps [zadd_commute, zminus_type, zadd_zminus_inverse]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
287 |
qed "zadd_zminus_inverse2"; |
| 0 | 288 |
|
| 438 | 289 |
goal Integ.thy "!!z. z:integ ==> z $+ $#0 = z"; |
| 0 | 290 |
by (rtac (zadd_commute RS trans) 1); |
| 438 | 291 |
by (REPEAT (ares_tac [znat_type, nat_0I, zadd_0] 1)); |
| 760 | 292 |
qed "zadd_0_right"; |
| 0 | 293 |
|
294 |
||
295 |
(*Need properties of $- ??? Or use $- just as an abbreviation? |
|
296 |
[| m: nat; n: nat; m>=n |] ==> $# (m #- n) = ($#m) $- ($#n) |
|
297 |
*) |
|
298 |
||
299 |
(**** zmult: multiplication on integ ****) |
|
300 |
||
301 |
(** Congruence property for multiplication **) |
|
302 |
||
303 |
goal Integ.thy |
|
| 1461 | 304 |
"congruent2(intrel, %p1 p2. \ |
305 |
\ split(%x1 y1. split(%x2 y2. \ |
|
| 0 | 306 |
\ intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1))";
|
307 |
by (rtac (equiv_intrel RS congruent2_commuteI) 1); |
|
| 4152 | 308 |
by Safe_tac; |
| 2469 | 309 |
by (ALLGOALS Asm_simp_tac); |
| 0 | 310 |
(*Proof that zmult is congruent in one argument*) |
| 438 | 311 |
by (asm_simp_tac |
| 4091 | 312 |
(simpset() addsimps (add_ac @ [add_mult_distrib_left RS sym])) 2); |
| 438 | 313 |
by (asm_simp_tac |
| 4091 | 314 |
(simpset() addsimps ([add_assoc RS sym, add_mult_distrib_left RS sym])) 2); |
| 0 | 315 |
(*Proof that zmult is commutative on representatives*) |
| 4091 | 316 |
by (asm_simp_tac (simpset() addsimps (mult_ac@add_ac)) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
317 |
qed "zmult_congruent2"; |
| 0 | 318 |
|
| 438 | 319 |
|
| 0 | 320 |
(*Resolve th against the corresponding facts for zmult*) |
321 |
val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2]; |
|
322 |
||
| 438 | 323 |
goalw Integ.thy [integ_def,zmult_def] |
324 |
"!!z w. [| z: integ; w: integ |] ==> z $* w : integ"; |
|
325 |
by (REPEAT (ares_tac [zmult_ize UN_equiv_class_type2, |
|
| 1461 | 326 |
split_type, add_type, mult_type, |
327 |
quotientI, SigmaI] 1)); |
|
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
328 |
qed "zmult_type"; |
| 0 | 329 |
|
| 438 | 330 |
goalw Integ.thy [zmult_def] |
| 1461 | 331 |
"!!x1 x2. [| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ |
332 |
\ (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) = \
|
|
| 438 | 333 |
\ intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}";
|
| 4091 | 334 |
by (asm_simp_tac (simpset() addsimps [zmult_ize UN_equiv_class2, SigmaI]) 1); |
| 760 | 335 |
qed "zmult"; |
| 0 | 336 |
|
337 |
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $* z = $#0"; |
|
| 438 | 338 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 339 |
by (asm_simp_tac (simpset() addsimps [zmult]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
340 |
qed "zmult_0"; |
| 0 | 341 |
|
|
16
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
7
diff
changeset
|
342 |
goalw Integ.thy [integ_def,znat_def] |
| 0 | 343 |
"!!z. z : integ ==> $#1 $* z = z"; |
| 438 | 344 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 345 |
by (asm_simp_tac (simpset() addsimps [zmult, add_0_right]) 1); |
| 760 | 346 |
qed "zmult_1"; |
| 0 | 347 |
|
348 |
goalw Integ.thy [integ_def] |
|
349 |
"!!z w. [| z: integ; w: integ |] ==> ($~ z) $* w = $~ (z $* w)"; |
|
| 438 | 350 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 351 |
by (asm_simp_tac (simpset() addsimps ([zminus, zmult] @ add_ac)) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
352 |
qed "zmult_zminus"; |
| 0 | 353 |
|
354 |
goalw Integ.thy [integ_def] |
|
355 |
"!!z w. [| z: integ; w: integ |] ==> ($~ z) $* ($~ w) = (z $* w)"; |
|
| 438 | 356 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 357 |
by (asm_simp_tac (simpset() addsimps ([zminus, zmult] @ add_ac)) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
358 |
qed "zmult_zminus_zminus"; |
| 0 | 359 |
|
360 |
goalw Integ.thy [integ_def] |
|
361 |
"!!z w. [| z: integ; w: integ |] ==> z $* w = w $* z"; |
|
| 438 | 362 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 363 |
by (asm_simp_tac (simpset() addsimps ([zmult] @ add_ac @ mult_ac)) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
364 |
qed "zmult_commute"; |
| 0 | 365 |
|
366 |
goalw Integ.thy [integ_def] |
|
367 |
"!!z1 z2 z3. [| z1: integ; z2: integ; z3: integ |] ==> \ |
|
368 |
\ (z1 $* z2) $* z3 = z1 $* (z2 $* z3)"; |
|
| 438 | 369 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
370 |
by (asm_simp_tac |
|
| 4091 | 371 |
(simpset() addsimps ([zmult, add_mult_distrib_left, |
| 1461 | 372 |
add_mult_distrib] @ add_ac @ mult_ac)) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
760
diff
changeset
|
373 |
qed "zmult_assoc"; |
| 0 | 374 |
|
|
904
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
375 |
(*For AC rewriting*) |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
376 |
qed_goal "zmult_left_commute" Integ.thy |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
377 |
"!!z1 z2 z3. [| z1:integ; z2:integ; z3: integ |] ==> \ |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
378 |
\ z1$*(z2$*z3) = z2$*(z1$*z3)" |
| 4091 | 379 |
(fn _ => [asm_simp_tac (simpset() addsimps [zmult_assoc RS sym, |
| 1461 | 380 |
zmult_commute]) 1]); |
|
904
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
381 |
|
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
382 |
(*Integer multiplication is an AC operator*) |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
383 |
val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute]; |
|
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
384 |
|
| 0 | 385 |
goalw Integ.thy [integ_def] |
386 |
"!!z1 z2 z3. [| z1: integ; z2: integ; w: integ |] ==> \ |
|
387 |
\ (z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)"; |
|
| 438 | 388 |
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); |
| 4091 | 389 |
by (asm_simp_tac (simpset() addsimps [zadd, zmult, add_mult_distrib]) 1); |
390 |
by (asm_simp_tac (simpset() addsimps (add_ac @ mult_ac)) 1); |
|
| 760 | 391 |
qed "zadd_zmult_distrib"; |
| 0 | 392 |
|
393 |
val integ_typechecks = |
|
394 |
[znat_type, zminus_type, zmagnitude_type, zadd_type, zmult_type]; |
|
395 |
||
| 2469 | 396 |
Addsimps integ_typechecks; |
|
904
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
397 |
|
| 2469 | 398 |
Addsimps [zadd_0, zadd_0_right, zadd_zminus_inverse, zadd_zminus_inverse2]; |
|
904
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
399 |
|
| 2469 | 400 |
Addsimps [zminus_zminus, zminus_0]; |
|
904
5240dcd12adb
Proved zadd_left_commute and zmult_left_commute to define
lcp
parents:
804
diff
changeset
|
401 |
|
| 2469 | 402 |
Addsimps [zmult_0, zmult_1, zmult_zminus]; |
403 |
||
404 |
Addsimps [zmagnitude_znat, zmagnitude_zminus_znat]; |