| author | nipkow | 
| Thu, 20 Dec 2001 18:22:44 +0100 | |
| changeset 12566 | fe20540bcf93 | 
| parent 11868 | 56db9f3a6b3e | 
| child 13438 | 527811f00c56 | 
| permissions | -rw-r--r-- | 
| 5508 | 1 | (* Title: IntDef.ML | 
| 2 | ID: $Id$ | |
| 3 | Authors: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1993 University of Cambridge | |
| 5 | ||
| 6 | The integers as equivalence classes over nat*nat. | |
| 7 | *) | |
| 8 | ||
| 9 | ||
| 11655 | 10 | Goalw [intrel_def] "(((x1,y1),(x2,y2)): intrel) = (x1+y2 = x2+y1)"; | 
| 9392 | 11 | by (Blast_tac 1); | 
| 5508 | 12 | qed "intrel_iff"; | 
| 13 | ||
| 9392 | 14 | Goalw [intrel_def, equiv_def, refl_def, sym_def, trans_def] | 
| 7375 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 15 | "equiv UNIV intrel"; | 
| 9392 | 16 | by Auto_tac; | 
| 5508 | 17 | qed "equiv_intrel"; | 
| 18 | ||
| 9392 | 19 | bind_thm ("equiv_intrel_iff", 
 | 
| 20 | [equiv_intrel, UNIV_I, UNIV_I] MRS eq_equiv_class_iff); | |
| 5508 | 21 | |
| 9392 | 22 | Goalw [Integ_def,intrel_def,quotient_def] | 
| 10834 | 23 |      "intrel``{(x,y)}:Integ";
 | 
| 5508 | 24 | by (Fast_tac 1); | 
| 25 | qed "intrel_in_integ"; | |
| 26 | ||
| 27 | Goal "inj_on Abs_Integ Integ"; | |
| 28 | by (rtac inj_on_inverseI 1); | |
| 29 | by (etac Abs_Integ_inverse 1); | |
| 30 | qed "inj_on_Abs_Integ"; | |
| 31 | ||
| 32 | Addsimps [equiv_intrel_iff, inj_on_Abs_Integ RS inj_on_iff, | |
| 33 | intrel_iff, intrel_in_integ, Abs_Integ_inverse]; | |
| 34 | ||
| 35 | Goal "inj(Rep_Integ)"; | |
| 36 | by (rtac inj_inverseI 1); | |
| 37 | by (rtac Rep_Integ_inverse 1); | |
| 38 | qed "inj_Rep_Integ"; | |
| 39 | ||
| 40 | ||
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5540diff
changeset | 41 | (** int: the injection from "nat" to "int" **) | 
| 5508 | 42 | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5540diff
changeset | 43 | Goal "inj int"; | 
| 5508 | 44 | by (rtac injI 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5540diff
changeset | 45 | by (rewtac int_def); | 
| 5508 | 46 | by (dtac (inj_on_Abs_Integ RS inj_onD) 1); | 
| 47 | by (REPEAT (rtac intrel_in_integ 1)); | |
| 48 | by (dtac eq_equiv_class 1); | |
| 49 | by (rtac equiv_intrel 1); | |
| 50 | by (Fast_tac 1); | |
| 9392 | 51 | by (asm_full_simp_tac (simpset() addsimps [intrel_def]) 1); | 
| 6991 | 52 | qed "inj_int"; | 
| 5508 | 53 | |
| 54 | ||
| 55 | (**** zminus: unary negation on Integ ****) | |
| 56 | ||
| 9392 | 57 | Goalw [congruent_def, intrel_def] | 
| 10834 | 58 |      "congruent intrel (%(x,y). intrel``{(y,x)})";
 | 
| 9392 | 59 | by (auto_tac (claset(), simpset() addsimps add_ac)); | 
| 5508 | 60 | qed "zminus_congruent"; | 
| 61 | ||
| 62 | Goalw [zminus_def] | |
| 10834 | 63 |       "- Abs_Integ(intrel``{(x,y)}) = Abs_Integ(intrel `` {(y,x)})";
 | 
| 5508 | 64 | by (simp_tac (simpset() addsimps | 
| 7375 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 65 | [equiv_intrel RS UN_equiv_class, zminus_congruent]) 1); | 
| 5508 | 66 | qed "zminus"; | 
| 67 | ||
| 7375 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 68 | (*Every integer can be written in the form Abs_Integ(...) *) | 
| 10834 | 69 | val [prem] = Goal "(!!x y. z = Abs_Integ(intrel``{(x,y)}) ==> P) ==> P";
 | 
| 5508 | 70 | by (res_inst_tac [("x1","z")] 
 | 
| 71 | (rewrite_rule [Integ_def] Rep_Integ RS quotientE) 1); | |
| 72 | by (dres_inst_tac [("f","Abs_Integ")] arg_cong 1);
 | |
| 73 | by (res_inst_tac [("p","x")] PairE 1);
 | |
| 74 | by (rtac prem 1); | |
| 75 | by (asm_full_simp_tac (simpset() addsimps [Rep_Integ_inverse]) 1); | |
| 76 | qed "eq_Abs_Integ"; | |
| 77 | ||
| 78 | Goal "- (- z) = (z::int)"; | |
| 79 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | |
| 80 | by (asm_simp_tac (simpset() addsimps [zminus]) 1); | |
| 81 | qed "zminus_zminus"; | |
| 82 | Addsimps [zminus_zminus]; | |
| 83 | ||
| 5594 | 84 | Goal "inj(%z::int. -z)"; | 
| 5508 | 85 | by (rtac injI 1); | 
| 86 | by (dres_inst_tac [("f","uminus")] arg_cong 1);
 | |
| 87 | by (Asm_full_simp_tac 1); | |
| 88 | qed "inj_zminus"; | |
| 89 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 90 | Goalw [int_def, Zero_int_def] "- 0 = (0::int)"; | 
| 5508 | 91 | by (simp_tac (simpset() addsimps [zminus]) 1); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 92 | qed "zminus_0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 93 | Addsimps [zminus_0]; | 
| 5508 | 94 | |
| 95 | ||
| 5540 | 96 | (**** neg: the test for negative integers ****) | 
| 5508 | 97 | |
| 98 | ||
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 99 | Goalw [neg_def, int_def] "~ neg(int n)"; | 
| 5508 | 100 | by (Simp_tac 1); | 
| 7010 
63120b6dca50
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6991diff
changeset | 101 | qed "not_neg_int"; | 
| 5508 | 102 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 103 | Goalw [neg_def, int_def] "neg(- (int (Suc n)))"; | 
| 5508 | 104 | by (simp_tac (simpset() addsimps [zminus]) 1); | 
| 7010 
63120b6dca50
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6991diff
changeset | 105 | qed "neg_zminus_int"; | 
| 5508 | 106 | |
| 7010 
63120b6dca50
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6991diff
changeset | 107 | Addsimps [neg_zminus_int, not_neg_int]; | 
| 5508 | 108 | |
| 109 | ||
| 110 | (**** zadd: addition on Integ ****) | |
| 111 | ||
| 112 | Goalw [zadd_def] | |
| 10834 | 113 |   "Abs_Integ(intrel``{(x1,y1)}) + Abs_Integ(intrel``{(x2,y2)}) = \
 | 
| 114 | \  Abs_Integ(intrel``{(x1+x2, y1+y2)})";
 | |
| 7375 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 115 | by (asm_simp_tac (simpset() addsimps [UN_UN_split_split_eq]) 1); | 
| 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 116 | by (stac (equiv_intrel RS UN_equiv_class2) 1); | 
| 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 117 | by (auto_tac (claset(), simpset() addsimps [congruent2_def])); | 
| 5508 | 118 | qed "zadd"; | 
| 119 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7010diff
changeset | 120 | Goal "- (z + w) = (- z) + (- w::int)"; | 
| 5508 | 121 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 122 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | |
| 123 | by (asm_simp_tac (simpset() addsimps [zminus,zadd]) 1); | |
| 124 | qed "zminus_zadd_distrib"; | |
| 125 | Addsimps [zminus_zadd_distrib]; | |
| 126 | ||
| 127 | Goal "(z::int) + w = w + z"; | |
| 128 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | |
| 129 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | |
| 5540 | 130 | by (asm_simp_tac (simpset() addsimps add_ac @ [zadd]) 1); | 
| 5508 | 131 | qed "zadd_commute"; | 
| 132 | ||
| 133 | Goal "((z1::int) + z2) + z3 = z1 + (z2 + z3)"; | |
| 134 | by (res_inst_tac [("z","z1")] eq_Abs_Integ 1);
 | |
| 135 | by (res_inst_tac [("z","z2")] eq_Abs_Integ 1);
 | |
| 136 | by (res_inst_tac [("z","z3")] eq_Abs_Integ 1);
 | |
| 137 | by (asm_simp_tac (simpset() addsimps [zadd, add_assoc]) 1); | |
| 138 | qed "zadd_assoc"; | |
| 139 | ||
| 140 | (*For AC rewriting*) | |
| 141 | Goal "(x::int)+(y+z)=y+(x+z)"; | |
| 142 | by (rtac (zadd_commute RS trans) 1); | |
| 143 | by (rtac (zadd_assoc RS trans) 1); | |
| 144 | by (rtac (zadd_commute RS arg_cong) 1); | |
| 145 | qed "zadd_left_commute"; | |
| 146 | ||
| 147 | (*Integer addition is an AC operator*) | |
| 7428 | 148 | bind_thms ("zadd_ac", [zadd_assoc,zadd_commute,zadd_left_commute]);
 | 
| 5508 | 149 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 150 | Goalw [int_def] "(int m) + (int n) = int (m + n)"; | 
| 5508 | 151 | by (simp_tac (simpset() addsimps [zadd]) 1); | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 152 | qed "zadd_int"; | 
| 5508 | 153 | |
| 5594 | 154 | Goal "(int m) + (int n + z) = int (m + n) + z"; | 
| 155 | by (simp_tac (simpset() addsimps [zadd_int, zadd_assoc RS sym]) 1); | |
| 156 | qed "zadd_int_left"; | |
| 157 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 158 | Goal "int (Suc m) = 1 + (int m)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 159 | by (simp_tac (simpset() addsimps [One_int_def, zadd_int]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 160 | qed "int_Suc"; | 
| 5508 | 161 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 162 | (*also for the instance declaration int :: plus_ac0*) | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 163 | Goalw [Zero_int_def, int_def] "(0::int) + z = z"; | 
| 5508 | 164 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 165 | by (asm_simp_tac (simpset() addsimps [zadd]) 1); | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 166 | qed "zadd_0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 167 | Addsimps [zadd_0]; | 
| 5508 | 168 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 169 | Goal "z + (0::int) = z"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 170 | by (rtac ([zadd_commute, zadd_0] MRS trans) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 171 | qed "zadd_0_right"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 172 | Addsimps [zadd_0_right]; | 
| 5508 | 173 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 174 | Goalw [int_def, Zero_int_def] "z + (- z) = (0::int)"; | 
| 5508 | 175 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 176 | by (asm_simp_tac (simpset() addsimps [zminus, zadd, add_commute]) 1); | |
| 5594 | 177 | qed "zadd_zminus_inverse"; | 
| 5508 | 178 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 179 | Goal "(- z) + z = (0::int)"; | 
| 5508 | 180 | by (rtac (zadd_commute RS trans) 1); | 
| 5594 | 181 | by (rtac zadd_zminus_inverse 1); | 
| 182 | qed "zadd_zminus_inverse2"; | |
| 5508 | 183 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 184 | Addsimps [zadd_zminus_inverse, zadd_zminus_inverse2]; | 
| 8949 | 185 | |
| 5508 | 186 | Goal "z + (- z + w) = (w::int)"; | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 187 | by (simp_tac (simpset() addsimps [zadd_assoc RS sym, zadd_0]) 1); | 
| 5508 | 188 | qed "zadd_zminus_cancel"; | 
| 189 | ||
| 190 | Goal "(-z) + (z + w) = (w::int)"; | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 191 | by (simp_tac (simpset() addsimps [zadd_assoc RS sym, zadd_0]) 1); | 
| 5508 | 192 | qed "zminus_zadd_cancel"; | 
| 193 | ||
| 194 | Addsimps [zadd_zminus_cancel, zminus_zadd_cancel]; | |
| 195 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 196 | Goal "(0::int) - x = -x"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 197 | by (simp_tac (simpset() addsimps [zdiff_def]) 1); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 198 | qed "zdiff0"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 199 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 200 | Goal "x - (0::int) = x"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 201 | by (simp_tac (simpset() addsimps [zdiff_def]) 1); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 202 | qed "zdiff0_right"; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 203 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 204 | Goal "x - x = (0::int)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 205 | by (simp_tac (simpset() addsimps [zdiff_def, Zero_int_def]) 1); | 
| 5594 | 206 | qed "zdiff_self"; | 
| 207 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 208 | Addsimps [zdiff0, zdiff0_right, zdiff_self]; | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 209 | |
| 5508 | 210 | |
| 211 | (** Lemmas **) | |
| 212 | ||
| 213 | Goal "(z::int) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"; | |
| 214 | by (asm_simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1); | |
| 215 | qed "zadd_assoc_cong"; | |
| 216 | ||
| 217 | Goal "(z::int) + (v + w) = v + (z + w)"; | |
| 218 | by (REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1)); | |
| 219 | qed "zadd_assoc_swap"; | |
| 220 | ||
| 221 | ||
| 222 | (**** zmult: multiplication on Integ ****) | |
| 223 | ||
| 7375 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 224 | (*Congruence property for multiplication*) | 
| 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 225 | Goal "congruent2 intrel \ | 
| 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 226 | \ (%p1 p2. (%(x1,y1). (%(x2,y2). \ | 
| 10834 | 227 | \                   intrel``{(x1*x2 + y1*y2, x1*y2 + y1*x2)}) p2) p1)";
 | 
| 5508 | 228 | by (rtac (equiv_intrel RS congruent2_commuteI) 1); | 
| 229 | by (pair_tac "w" 2); | |
| 7375 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 230 | by (ALLGOALS Clarify_tac); | 
| 5508 | 231 | by (simp_tac (simpset() addsimps add_ac@mult_ac) 1); | 
| 232 | by (asm_simp_tac (simpset() delsimps [equiv_intrel_iff] | |
| 7375 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 233 | addsimps add_ac@mult_ac) 1); | 
| 9392 | 234 | by (rename_tac "x1 x2 y1 y2 z1 z2" 1); | 
| 235 | by (rtac ([equiv_intrel, intrel_iff RS iffD2] MRS equiv_class_eq) 1); | |
| 236 | by (asm_full_simp_tac (simpset() addsimps [intrel_def]) 1); | |
| 237 | by (subgoal_tac | |
| 238 | "x1*z1 + y2*z1 = y1*z1 + x2*z1 & x1*z2 + y2*z2 = y1*z2 + x2*z2" 1); | |
| 239 | by (asm_simp_tac (simpset() addsimps [add_mult_distrib RS sym]) 2); | |
| 240 | by (arith_tac 1); | |
| 5508 | 241 | qed "zmult_congruent2"; | 
| 242 | ||
| 243 | Goalw [zmult_def] | |
| 10834 | 244 |    "Abs_Integ((intrel``{(x1,y1)})) * Abs_Integ((intrel``{(x2,y2)})) =   \
 | 
| 245 | \   Abs_Integ(intrel `` {(x1*x2 + y1*y2, x1*y2 + y1*x2)})";
 | |
| 7375 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 246 | by (asm_simp_tac | 
| 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 247 | (simpset() addsimps [UN_UN_split_split_eq, zmult_congruent2, | 
| 
2cb340e66d15
tidied, allowing pattern-matching in defs of zadd and zmult
 paulson parents: 
7127diff
changeset | 248 | equiv_intrel RS UN_equiv_class2]) 1); | 
| 5508 | 249 | qed "zmult"; | 
| 250 | ||
| 251 | Goal "(- z) * w = - (z * (w::int))"; | |
| 252 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | |
| 253 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | |
| 5540 | 254 | by (asm_simp_tac (simpset() addsimps [zminus, zmult] @ add_ac) 1); | 
| 5508 | 255 | qed "zmult_zminus"; | 
| 256 | ||
| 257 | Goal "(z::int) * w = w * z"; | |
| 258 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | |
| 259 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | |
| 5540 | 260 | by (asm_simp_tac (simpset() addsimps [zmult] @ add_ac @ mult_ac) 1); | 
| 5508 | 261 | qed "zmult_commute"; | 
| 262 | ||
| 263 | Goal "((z1::int) * z2) * z3 = z1 * (z2 * z3)"; | |
| 264 | by (res_inst_tac [("z","z1")] eq_Abs_Integ 1);
 | |
| 265 | by (res_inst_tac [("z","z2")] eq_Abs_Integ 1);
 | |
| 266 | by (res_inst_tac [("z","z3")] eq_Abs_Integ 1);
 | |
| 5540 | 267 | by (asm_simp_tac (simpset() addsimps [add_mult_distrib2,zmult] @ | 
| 268 | add_ac @ mult_ac) 1); | |
| 5508 | 269 | qed "zmult_assoc"; | 
| 270 | ||
| 271 | (*For AC rewriting*) | |
| 272 | Goal "(z1::int)*(z2*z3) = z2*(z1*z3)"; | |
| 273 | by (rtac (zmult_commute RS trans) 1); | |
| 274 | by (rtac (zmult_assoc RS trans) 1); | |
| 275 | by (rtac (zmult_commute RS arg_cong) 1); | |
| 276 | qed "zmult_left_commute"; | |
| 277 | ||
| 278 | (*Integer multiplication is an AC operator*) | |
| 7428 | 279 | bind_thms ("zmult_ac", [zmult_assoc, zmult_commute, zmult_left_commute]);
 | 
| 5508 | 280 | |
| 281 | Goal "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"; | |
| 282 | by (res_inst_tac [("z","z1")] eq_Abs_Integ 1);
 | |
| 283 | by (res_inst_tac [("z","z2")] eq_Abs_Integ 1);
 | |
| 284 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | |
| 285 | by (asm_simp_tac | |
| 5540 | 286 | (simpset() addsimps [add_mult_distrib2, zadd, zmult] @ | 
| 287 | add_ac @ mult_ac) 1); | |
| 5508 | 288 | qed "zadd_zmult_distrib"; | 
| 289 | ||
| 9544 
f9202e219a29
added a dummy "thm list" argument to prove_conv for the new interface to
 paulson parents: 
9392diff
changeset | 290 | val zmult_commute'= inst "z" "w" zmult_commute; | 
| 5508 | 291 | |
| 292 | Goal "w * (- z) = - (w * (z::int))"; | |
| 293 | by (simp_tac (simpset() addsimps [zmult_commute', zmult_zminus]) 1); | |
| 294 | qed "zmult_zminus_right"; | |
| 295 | ||
| 296 | Goal "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"; | |
| 297 | by (simp_tac (simpset() addsimps [zmult_commute',zadd_zmult_distrib]) 1); | |
| 298 | qed "zadd_zmult_distrib2"; | |
| 299 | ||
| 6839 | 300 | Goalw [zdiff_def] "((z1::int) - z2) * w = (z1 * w) - (z2 * w)"; | 
| 301 | by (stac zadd_zmult_distrib 1); | |
| 302 | by (simp_tac (simpset() addsimps [zmult_zminus]) 1); | |
| 303 | qed "zdiff_zmult_distrib"; | |
| 304 | ||
| 305 | Goal "(w::int) * (z1 - z2) = (w * z1) - (w * z2)"; | |
| 306 | by (simp_tac (simpset() addsimps [zmult_commute',zdiff_zmult_distrib]) 1); | |
| 307 | qed "zdiff_zmult_distrib2"; | |
| 308 | ||
| 10451 | 309 | bind_thms ("int_distrib",
 | 
| 310 | [zadd_zmult_distrib, zadd_zmult_distrib2, zdiff_zmult_distrib, zdiff_zmult_distrib2]); | |
| 311 | ||
| 7010 
63120b6dca50
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6991diff
changeset | 312 | Goalw [int_def] "(int m) * (int n) = int (m * n)"; | 
| 
63120b6dca50
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6991diff
changeset | 313 | by (simp_tac (simpset() addsimps [zmult]) 1); | 
| 
63120b6dca50
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6991diff
changeset | 314 | qed "zmult_int"; | 
| 
63120b6dca50
more renaming of theorems from _nat to _int (corresponding to a function that
 paulson parents: 
6991diff
changeset | 315 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 316 | Goalw [Zero_int_def, int_def] "0 * z = (0::int)"; | 
| 5508 | 317 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 318 | by (asm_simp_tac (simpset() addsimps [zmult]) 1); | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 319 | qed "zmult_0"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 320 | Addsimps [zmult_0]; | 
| 5508 | 321 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 322 | Goalw [One_int_def, int_def] "(1::int) * z = z"; | 
| 5508 | 323 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 324 | by (asm_simp_tac (simpset() addsimps [zmult]) 1); | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 325 | qed "zmult_1"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 326 | Addsimps [zmult_1]; | 
| 5508 | 327 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 328 | Goal "z * 0 = (0::int)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 329 | by (rtac ([zmult_commute, zmult_0] MRS trans) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 330 | qed "zmult_0_right"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 331 | Addsimps [zmult_0_right]; | 
| 5508 | 332 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 333 | Goal "z * (1::int) = z"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 334 | by (rtac ([zmult_commute, zmult_1] MRS trans) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 335 | qed "zmult_1_right"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 336 | Addsimps [zmult_1_right]; | 
| 5508 | 337 | |
| 338 | ||
| 339 | (* Theorems about less and less_equal *) | |
| 340 | ||
| 341 | (*This lemma allows direct proofs of other <-properties*) | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5540diff
changeset | 342 | Goalw [zless_def, neg_def, zdiff_def, int_def] | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 343 | "(w < z) = (EX n. z = w + int(Suc n))"; | 
| 5508 | 344 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 345 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | |
| 346 | by (Clarify_tac 1); | |
| 347 | by (asm_full_simp_tac (simpset() addsimps [zadd, zminus]) 1); | |
| 10558 | 348 | by (safe_tac (claset() addSDs [less_imp_Suc_add])); | 
| 5508 | 349 | by (res_inst_tac [("x","k")] exI 1);
 | 
| 350 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps add_ac))); | |
| 351 | qed "zless_iff_Suc_zadd"; | |
| 352 | ||
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 353 | Goal "z < z + int (Suc n)"; | 
| 5508 | 354 | by (auto_tac (claset(), | 
| 355 | simpset() addsimps [zless_iff_Suc_zadd, zadd_assoc, | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 356 | zadd_int])); | 
| 5508 | 357 | qed "zless_zadd_Suc"; | 
| 358 | ||
| 359 | Goal "[| z1<z2; z2<z3 |] ==> z1 < (z3::int)"; | |
| 360 | by (auto_tac (claset(), | |
| 361 | simpset() addsimps [zless_iff_Suc_zadd, zadd_assoc, | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 362 | zadd_int])); | 
| 5508 | 363 | qed "zless_trans"; | 
| 364 | ||
| 365 | Goal "!!w::int. z<w ==> ~w<z"; | |
| 366 | by (safe_tac (claset() addSDs [zless_iff_Suc_zadd RS iffD1])); | |
| 367 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | |
| 368 | by Safe_tac; | |
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5540diff
changeset | 369 | by (asm_full_simp_tac (simpset() addsimps [int_def, zadd]) 1); | 
| 5508 | 370 | qed "zless_not_sym"; | 
| 371 | ||
| 372 | (* [| n<m; ~P ==> m<n |] ==> P *) | |
| 5540 | 373 | bind_thm ("zless_asym", zless_not_sym RS swap);
 | 
| 5508 | 374 | |
| 375 | Goal "!!z::int. ~ z<z"; | |
| 376 | by (resolve_tac [zless_asym RS notI] 1); | |
| 377 | by (REPEAT (assume_tac 1)); | |
| 378 | qed "zless_not_refl"; | |
| 379 | ||
| 380 | (* z<z ==> R *) | |
| 5594 | 381 | bind_thm ("zless_irrefl", zless_not_refl RS notE);
 | 
| 5508 | 382 | AddSEs [zless_irrefl]; | 
| 383 | ||
| 384 | ||
| 385 | (*"Less than" is a linear ordering*) | |
| 5540 | 386 | Goalw [zless_def, neg_def, zdiff_def] | 
| 5508 | 387 | "z<w | z=w | w<(z::int)"; | 
| 388 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | |
| 389 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | |
| 390 | by Safe_tac; | |
| 391 | by (asm_full_simp_tac | |
| 392 | (simpset() addsimps [zadd, zminus, Image_iff, Bex_def]) 1); | |
| 393 | by (res_inst_tac [("m1", "x+ya"), ("n1", "xa+y")] (less_linear RS disjE) 1);
 | |
| 5758 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5594diff
changeset | 394 | by (ALLGOALS (force_tac (claset(), simpset() addsimps add_ac))); | 
| 5508 | 395 | qed "zless_linear"; | 
| 396 | ||
| 397 | Goal "!!w::int. (w ~= z) = (w<z | z<w)"; | |
| 398 | by (cut_facts_tac [zless_linear] 1); | |
| 399 | by (Blast_tac 1); | |
| 400 | qed "int_neq_iff"; | |
| 401 | ||
| 402 | (*** eliminates ~= in premises ***) | |
| 403 | bind_thm("int_neqE", int_neq_iff RS iffD1 RS disjE);
 | |
| 404 | ||
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 405 | Goal "(int m = int n) = (m = n)"; | 
| 6991 | 406 | by (fast_tac (claset() addSEs [inj_int RS injD]) 1); | 
| 5562 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5540diff
changeset | 407 | qed "int_int_eq"; | 
| 
02261e6880d1
Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
 paulson parents: 
5540diff
changeset | 408 | AddIffs [int_int_eq]; | 
| 5508 | 409 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 410 | Goal "(int n = 0) = (n = 0)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 411 | by (simp_tac (simpset() addsimps [Zero_int_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 412 | qed "int_eq_0_conv"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 413 | Addsimps [int_eq_0_conv]; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 414 | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 415 | Goal "(int m < int n) = (m<n)"; | 
| 5508 | 416 | by (simp_tac (simpset() addsimps [less_iff_Suc_add, zless_iff_Suc_zadd, | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 417 | zadd_int]) 1); | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 418 | qed "zless_int"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 419 | Addsimps [zless_int]; | 
| 5508 | 420 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 421 | Goal "~ (int k < 0)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 422 | by (simp_tac (simpset() addsimps [Zero_int_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 423 | qed "int_less_0_conv"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 424 | Addsimps [int_less_0_conv]; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 425 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 426 | Goal "(0 < int n) = (0 < n)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 427 | by (simp_tac (simpset() addsimps [Zero_int_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 428 | qed "zero_less_int_conv"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 429 | Addsimps [zero_less_int_conv]; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 430 | |
| 5508 | 431 | |
| 432 | (*** Properties of <= ***) | |
| 433 | ||
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 434 | Goalw [zle_def, le_def] "(int m <= int n) = (m<=n)"; | 
| 5508 | 435 | by (Simp_tac 1); | 
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 436 | qed "zle_int"; | 
| 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 437 | Addsimps [zle_int]; | 
| 5508 | 438 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 439 | Goal "(0 <= int n)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 440 | by (simp_tac (simpset() addsimps [Zero_int_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 441 | qed "zero_zle_int"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 442 | Addsimps [zero_zle_int]; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 443 | |
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 444 | Goal "(int n <= 0) = (n = 0)"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 445 | by (simp_tac (simpset() addsimps [Zero_int_def]) 1); | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 446 | qed "int_le_0_conv"; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 447 | Addsimps [int_le_0_conv]; | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 448 | |
| 5508 | 449 | Goalw [zle_def] "z <= w ==> z < w | z=(w::int)"; | 
| 450 | by (cut_facts_tac [zless_linear] 1); | |
| 451 | by (blast_tac (claset() addEs [zless_asym]) 1); | |
| 452 | qed "zle_imp_zless_or_eq"; | |
| 453 | ||
| 454 | Goalw [zle_def] "z<w | z=w ==> z <= (w::int)"; | |
| 455 | by (cut_facts_tac [zless_linear] 1); | |
| 456 | by (blast_tac (claset() addEs [zless_asym]) 1); | |
| 457 | qed "zless_or_eq_imp_zle"; | |
| 458 | ||
| 459 | Goal "(x <= (y::int)) = (x < y | x=y)"; | |
| 460 | by (REPEAT(ares_tac [iffI, zless_or_eq_imp_zle, zle_imp_zless_or_eq] 1)); | |
| 10472 | 461 | qed "int_le_less"; | 
| 5508 | 462 | |
| 463 | Goal "w <= (w::int)"; | |
| 10472 | 464 | by (simp_tac (simpset() addsimps [int_le_less]) 1); | 
| 5508 | 465 | qed "zle_refl"; | 
| 466 | ||
| 467 | (* Axiom 'linorder_linear' of class 'linorder': *) | |
| 468 | Goal "(z::int) <= w | w <= z"; | |
| 10472 | 469 | by (simp_tac (simpset() addsimps [int_le_less]) 1); | 
| 5508 | 470 | by (cut_facts_tac [zless_linear] 1); | 
| 471 | by (Blast_tac 1); | |
| 472 | qed "zle_linear"; | |
| 473 | ||
| 10647 | 474 | (* Axiom 'order_trans of class 'order': *) | 
| 5508 | 475 | Goal "[| i <= j; j <= k |] ==> i <= (k::int)"; | 
| 476 | by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq, | |
| 477 | rtac zless_or_eq_imp_zle, | |
| 478 | blast_tac (claset() addIs [zless_trans])]); | |
| 479 | qed "zle_trans"; | |
| 480 | ||
| 481 | Goal "[| z <= w; w <= z |] ==> z = (w::int)"; | |
| 482 | by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq, | |
| 483 | blast_tac (claset() addEs [zless_asym])]); | |
| 484 | qed "zle_anti_sym"; | |
| 485 | ||
| 486 | (* Axiom 'order_less_le' of class 'order': *) | |
| 11655 | 487 | Goal "((w::int) < z) = (w <= z & w ~= z)"; | 
| 5508 | 488 | by (simp_tac (simpset() addsimps [zle_def, int_neq_iff]) 1); | 
| 489 | by (blast_tac (claset() addSEs [zless_asym]) 1); | |
| 490 | qed "int_less_le"; | |
| 491 | ||
| 492 | ||
| 493 | (*** Subtraction laws ***) | |
| 494 | ||
| 495 | Goal "x + (y - z) = (x + y) - (z::int)"; | |
| 5540 | 496 | by (simp_tac (simpset() addsimps zdiff_def::zadd_ac) 1); | 
| 5508 | 497 | qed "zadd_zdiff_eq"; | 
| 498 | ||
| 499 | Goal "(x - y) + z = (x + z) - (y::int)"; | |
| 5540 | 500 | by (simp_tac (simpset() addsimps zdiff_def::zadd_ac) 1); | 
| 5508 | 501 | qed "zdiff_zadd_eq"; | 
| 502 | ||
| 503 | Goal "(x - y) - z = x - (y + (z::int))"; | |
| 5540 | 504 | by (simp_tac (simpset() addsimps zdiff_def::zadd_ac) 1); | 
| 5508 | 505 | qed "zdiff_zdiff_eq"; | 
| 506 | ||
| 507 | Goal "x - (y - z) = (x + z) - (y::int)"; | |
| 5540 | 508 | by (simp_tac (simpset() addsimps zdiff_def::zadd_ac) 1); | 
| 5508 | 509 | qed "zdiff_zdiff_eq2"; | 
| 510 | ||
| 511 | Goalw [zless_def, zdiff_def] "(x-y < z) = (x < z + (y::int))"; | |
| 512 | by (simp_tac (simpset() addsimps zadd_ac) 1); | |
| 513 | qed "zdiff_zless_eq"; | |
| 514 | ||
| 515 | Goalw [zless_def, zdiff_def] "(x < z-y) = (x + (y::int) < z)"; | |
| 516 | by (simp_tac (simpset() addsimps zadd_ac) 1); | |
| 517 | qed "zless_zdiff_eq"; | |
| 518 | ||
| 519 | Goalw [zle_def] "(x-y <= z) = (x <= z + (y::int))"; | |
| 520 | by (simp_tac (simpset() addsimps [zless_zdiff_eq]) 1); | |
| 521 | qed "zdiff_zle_eq"; | |
| 522 | ||
| 523 | Goalw [zle_def] "(x <= z-y) = (x + (y::int) <= z)"; | |
| 524 | by (simp_tac (simpset() addsimps [zdiff_zless_eq]) 1); | |
| 525 | qed "zle_zdiff_eq"; | |
| 526 | ||
| 527 | Goalw [zdiff_def] "(x-y = z) = (x = z + (y::int))"; | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 528 | by (auto_tac (claset(), | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 529 | simpset() addsimps [zadd_assoc, symmetric Zero_int_def])); | 
| 5508 | 530 | qed "zdiff_eq_eq"; | 
| 531 | ||
| 532 | Goalw [zdiff_def] "(x = z-y) = (x + (y::int) = z)"; | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 533 | by (auto_tac (claset(), | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 534 | simpset() addsimps [zadd_assoc, symmetric Zero_int_def])); | 
| 5508 | 535 | qed "eq_zdiff_eq"; | 
| 536 | ||
| 537 | (*This list of rewrites simplifies (in)equalities by bringing subtractions | |
| 538 | to the top and then moving negative terms to the other side. | |
| 539 | Use with zadd_ac*) | |
| 9108 | 540 | bind_thms ("zcompare_rls",
 | 
| 5508 | 541 | [symmetric zdiff_def, | 
| 542 | zadd_zdiff_eq, zdiff_zadd_eq, zdiff_zdiff_eq, zdiff_zdiff_eq2, | |
| 543 | zdiff_zless_eq, zless_zdiff_eq, zdiff_zle_eq, zle_zdiff_eq, | |
| 9108 | 544 | zdiff_eq_eq, eq_zdiff_eq]); | 
| 5508 | 545 | |
| 546 | ||
| 547 | (** Cancellation laws **) | |
| 548 | ||
| 549 | Goal "!!w::int. (z + w' = z + w) = (w' = w)"; | |
| 550 | by Safe_tac; | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7010diff
changeset | 551 | by (dres_inst_tac [("f", "%x. x + (-z)")] arg_cong 1);
 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 552 | by (asm_full_simp_tac (simpset() addsimps symmetric Zero_int_def :: | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 553 | zadd_ac) 1); | 
| 5508 | 554 | qed "zadd_left_cancel"; | 
| 555 | ||
| 556 | Addsimps [zadd_left_cancel]; | |
| 557 | ||
| 558 | Goal "!!z::int. (w' + z = w + z) = (w' = w)"; | |
| 559 | by (asm_full_simp_tac (simpset() addsimps zadd_ac) 1); | |
| 560 | qed "zadd_right_cancel"; | |
| 561 | ||
| 562 | Addsimps [zadd_right_cancel]; | |
| 563 | ||
| 564 | ||
| 5594 | 565 | (** For the cancellation simproc. | 
| 566 | The idea is to cancel like terms on opposite sides by subtraction **) | |
| 567 | ||
| 568 | Goal "(x::int) - y = x' - y' ==> (x<y) = (x'<y')"; | |
| 569 | by (asm_simp_tac (simpset() addsimps [zless_def]) 1); | |
| 570 | qed "zless_eqI"; | |
| 5508 | 571 | |
| 5594 | 572 | Goal "(x::int) - y = x' - y' ==> (y<=x) = (y'<=x')"; | 
| 573 | by (dtac zless_eqI 1); | |
| 574 | by (asm_simp_tac (simpset() addsimps [zle_def]) 1); | |
| 575 | qed "zle_eqI"; | |
| 5508 | 576 | |
| 5594 | 577 | Goal "(x::int) - y = x' - y' ==> (x=y) = (x'=y')"; | 
| 578 | by Safe_tac; | |
| 579 | by (ALLGOALS | |
| 580 | (asm_full_simp_tac (simpset() addsimps [eq_zdiff_eq, zdiff_eq_eq]))); | |
| 581 | qed "zeq_eqI"; | |
| 5582 
a356fb49e69e
many renamings and changes.  Simproc for cancelling common terms in relations
 paulson parents: 
5562diff
changeset | 582 |