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